11fbee269b
To provide typing and auto-completion information, we create a background process by starting lean with the option --server. In "server" mode, we create snapshots of the parser state. The idea is to be able to quickly reprocess a file when the user makes a modification. This commit fixes a bug in the snapshot operation. It was not saving in the snapshots the unique name generator used by the parser. By creating a fresh name generator, we may accidentally assign the same internal name to two different constants. This bug triggered the crash described in issue #593. This commit fixes the test in the comment https://github.com/leanprover/lean/issues/593#issuecomment-101768484 This commit also adds a smaller version of the problem to the test suite
32 lines
790 B
Text
32 lines
790 B
Text
import algebra.group data.set
|
|
namespace group_hom
|
|
open algebra
|
|
-- ⁻¹ in eq.ops conflicts with group ⁻¹
|
|
-- open eq.ops
|
|
open set
|
|
open function
|
|
local attribute set [reducible]
|
|
structure is_subgroup [class] {A : Type} (H : set A) : Type
|
|
|
|
section
|
|
variables {A B : Type}
|
|
variable [s1 : group A]
|
|
variable [s2 : group B]
|
|
include s1
|
|
include s2
|
|
definition is_hom (f : A → B) := ∀ a b, f (a*b) = (f a)*(f b)
|
|
|
|
variable f : A → B
|
|
variable Hom : is_hom f
|
|
definition ker : set A := λ a, (f a) = 1
|
|
|
|
lemma ker.has_one (Hom : is_hom f) : 1 ∈ ker f := sorry
|
|
theorem hom_map_one : f 1 = 1 := sorry
|
|
theorem hom_map_mul_closed (Hom : is_hom f) (H : set A) : unit :=
|
|
sorry
|
|
|
|
variable {H : set A}
|
|
variable [is_subgH : is_subgroup H]
|
|
include is_subgH
|
|
end
|
|
end group_hom
|