diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 9551ba15a5008..3441ca351115d 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -638,43 +638,30 @@ lemma closure_induction {p : G → Prop} {x} (h : x ∈ closure k) (Hinv : ∀ x, p x → p x⁻¹) : p x := (@closure_le _ _ ⟨p, H1, Hmul, Hinv⟩ _).2 Hk h -/-- An induction principle on elements of the subtype `subgroup.closure`. -If `p` holds for `1` and all elements of `k`, and is preserved under multiplication and inverse, -then `p` holds for all elements `x : closure k`. - -The difference with `subgroup.closure_induction` is that this acts on the subtype. --/ -@[elab_as_eliminator, to_additive "An induction principle on elements of the subtype -`add_subgroup.closure`. If `p` holds for `0` and all elements of `k`, and is preserved under -addition and negation, then `p` holds for all elements `x : closure k`. - -The difference with `add_subgroup.closure_induction` is that this acts on the subtype."] -lemma closure_induction' (k : set G) {p : closure k → Prop} - (Hk : ∀ x (h : x ∈ k), p ⟨x, subset_closure h⟩) - (H1 : p 1) - (Hmul : ∀ x y, p x → p y → p (x * y)) - (Hinv : ∀ x, p x → p x⁻¹) - (x : closure k) : - p x := -subtype.rec_on x $ λ x hx, begin - refine exists.elim _ (λ (hx : x ∈ closure k) (hc : p ⟨x, hx⟩), hc), +/-- A dependent version of `subgroup.closure_induction`. -/ +@[elab_as_eliminator, to_additive "A dependent version of `add_subgroup.closure_induction`. "] +lemma closure_induction' {p : Π x, x ∈ closure k → Prop} + (Hs : ∀ x (h : x ∈ k), p x (subset_closure h)) + (H1 : p 1 (one_mem _)) + (Hmul : ∀ x hx y hy, p x hx → p y hy → p (x * y) (mul_mem _ hx hy)) + (Hinv : ∀ x hx, p x hx → p x⁻¹ (inv_mem _ hx)) + {x} (hx : x ∈ closure k) : + p x hx := +begin + refine exists.elim _ (λ (hx : x ∈ closure k) (hc : p x hx), hc), exact closure_induction hx - (λ x hx, ⟨subset_closure hx, Hk x hx⟩) - ⟨one_mem _, H1⟩ - (λ x y hx hy, exists.elim hx $ λ hx' hx, exists.elim hy $ λ hy' hy, - ⟨mul_mem _ hx' hy', Hmul _ _ hx hy⟩) - (λ x hx, exists.elim hx $ λ hx' hx, ⟨inv_mem _ hx', Hinv _ hx⟩), + (λ x hx, ⟨_, Hs x hx⟩) ⟨_, H1⟩ (λ x y ⟨hx', hx⟩ ⟨hy', hy⟩, ⟨_, Hmul _ _ _ _ hx hy⟩) + (λ x ⟨hx', hx⟩, ⟨_, Hinv _ _ hx⟩), end @[simp, to_additive] lemma closure_closure_coe_preimage {k : set G} : closure ((coe : closure k → G) ⁻¹' k) = ⊤ := -begin - refine eq_top_iff.2 (λ x hx, closure_induction' (λ x, _) _ _ (λ g₁ g₂ hg₁ hg₂, _) (λ g hg, _) x), - { intros g hg, - exact subset_closure hg }, - { exact subgroup.one_mem _ }, - { exact subgroup.mul_mem _ hg₁ hg₂ }, - { exact subgroup.inv_mem _ hg } +eq_top_iff.2 $ λ x, subtype.rec_on x $ λ x hx _, begin + refine closure_induction' (λ g hg, _) _ (λ g₁ g₂ hg₁ hg₂, _) (λ g hg, _) hx, + { exact subset_closure hg }, + { exact one_mem _ }, + { exact mul_mem _ }, + { exact inv_mem _ } end variable (G) diff --git a/src/group_theory/submonoid/basic.lean b/src/group_theory/submonoid/basic.lean index ae2d0a9a91532..fd578df11c822 100644 --- a/src/group_theory/submonoid/basic.lean +++ b/src/group_theory/submonoid/basic.lean @@ -275,6 +275,20 @@ lemma closure_induction {p : M → Prop} {x} (h : x ∈ closure s) (Hmul : ∀ x y, p x → p y → p (x * y)) : p x := (@closure_le _ _ _ ⟨p, H1, Hmul⟩).2 Hs h +/-- A dependent version of `submonoid.closure_induction`. -/ +@[elab_as_eliminator, to_additive "A dependent version of `add_submonoid.closure_induction`. "] +lemma closure_induction' (s : set M) {p : Π x, x ∈ closure s → Prop} + (Hs : ∀ x (h : x ∈ s), p x (subset_closure h)) + (H1 : p 1 (one_mem _)) + (Hmul : ∀ x hx y hy, p x hx → p y hy → p (x * y) (mul_mem _ hx hy)) + {x} (hx : x ∈ closure s) : + p x hx := +begin + refine exists.elim _ (λ (hx : x ∈ closure s) (hc : p x hx), hc), + exact closure_induction hx + (λ x hx, ⟨_, Hs x hx⟩) ⟨_, H1⟩ (λ x y ⟨hx', hx⟩ ⟨hy', hy⟩, ⟨_, Hmul _ _ _ _ hx hy⟩), +end + /-- If `s` is a dense set in a monoid `M`, `submonoid.closure s = ⊤`, then in order to prove that some predicate `p` holds for all `x : M` it suffices to verify `p x` for `x ∈ s`, verify `p 1`, and verify that `p x` and `p y` imply `p (x * y)`. -/ diff --git a/src/group_theory/submonoid/operations.lean b/src/group_theory/submonoid/operations.lean index e06743f5355c7..4767adc91898f 100644 --- a/src/group_theory/submonoid/operations.lean +++ b/src/group_theory/submonoid/operations.lean @@ -462,40 +462,13 @@ noncomputable def equiv_map_of_injective (f : M →* N) (hf : function.injective f) (x : S) : (equiv_map_of_injective S f hf x : N) = f x := rfl -/-- An induction principle on elements of the type `submonoid.closure s`. -If `p` holds for `1` and all elements of `s`, and is preserved under multiplication, then `p` -holds for all elements of the closure of `s`. - -The difference with `submonoid.closure_induction` is that this acts on the subtype. --/ -@[elab_as_eliminator, to_additive "An induction principle on elements of the type -`add_submonoid.closure s`. If `p` holds for `0` and all elements of `s`, and is preserved under -addition, then `p` holds for all elements of the closure of `s`. - -The difference with `add_submonoid.closure_induction` is that this acts on the subtype."] -lemma closure_induction' (s : set M) {p : closure s → Prop} - (Hs : ∀ x (h : x ∈ s), p ⟨x, subset_closure h⟩) - (H1 : p 1) - (Hmul : ∀ x y, p x → p y → p (x * y)) - (x : closure s) : - p x := -subtype.rec_on x $ λ x hx, begin - refine exists.elim _ (λ (hx : x ∈ closure s) (hc : p ⟨x, hx⟩), hc), - exact closure_induction hx - (λ x hx, ⟨subset_closure hx, Hs x hx⟩) - ⟨one_mem _, H1⟩ - (λ x y hx hy, exists.elim hx $ λ hx' hx, exists.elim hy $ λ hy' hy, - ⟨mul_mem _ hx' hy', Hmul _ _ hx hy⟩), -end - @[simp, to_additive] lemma closure_closure_coe_preimage {s : set M} : closure ((coe : closure s → M) ⁻¹' s) = ⊤ := -begin - refine eq_top_iff.2 (λ x hx, closure_induction' (λ x, _) _ _ (λ g₁ g₂ hg₁ hg₂, _) x), - { intros g hg, - exact subset_closure hg }, +eq_top_iff.2 $ λ x, subtype.rec_on x $ λ x hx _, begin + refine closure_induction' _ (λ g hg, _) _ (λ g₁ g₂ hg₁ hg₂, _) hx, + { exact subset_closure hg }, { exact submonoid.one_mem _ }, - { exact submonoid.mul_mem _ hg₁ hg₂ }, + { exact submonoid.mul_mem _ }, end /-- Given `submonoid`s `s`, `t` of monoids `M`, `N` respectively, `s × t` as a submonoid diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 6781e4b8bcab7..f9077158864ab 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -860,25 +860,28 @@ preserved under addition and scalar multiplication, then `p` holds for all eleme (H2 : ∀ (a:R) x, p x → p (a • x)) : p x := (@span_le _ _ _ _ _ _ ⟨p, H0, H1, H2⟩).2 Hs h -/-- The difference with `submodule.span_induction` is that this acts on the subtype. -/ -lemma span_induction' {p : span R s → Prop} (Hs : ∀ x (h : x ∈ s), p ⟨x, subset_span h⟩) (H0 : p 0) - (H1 : ∀ x y, p x → p y → p (x + y)) (H2 : ∀ (a : R) x, p x → p (a • x)) (x : span R s) : p x := -subtype.rec_on x $ λ x hx, begin - refine exists.elim _ (λ (hx : x ∈ span R s) (hc : p ⟨x, hx⟩), hc), +/-- A dependent version of `submodule.span_induction`. -/ +lemma span_induction' {p : Π x, x ∈ span R s → Prop} + (Hs : ∀ x (h : x ∈ s), p x (subset_span h)) + (H0 : p 0 (submodule.zero_mem _)) + (H1 : ∀ x hx y hy, p x hx → p y hy → p (x + y) (submodule.add_mem _ ‹_› ‹_›)) + (H2 : ∀ (a : R) x hx, p x hx → p (a • x) (submodule.smul_mem _ _ ‹_›)) {x} (hx : x ∈ span R s) : + p x hx := +begin + refine exists.elim _ (λ (hx : x ∈ span R s) (hc : p x hx), hc), refine span_induction hx (λ m hm, ⟨subset_span hm, Hs m hm⟩) ⟨zero_mem _, H0⟩ (λ x y hx hy, exists.elim hx $ λ hx' hx, exists.elim hy $ λ hy' hy, - ⟨add_mem _ hx' hy', H1 _ _ hx hy⟩) (λ r x hx, exists.elim hx $ λ hx' hx, - ⟨smul_mem _ _ hx', H2 r _ hx⟩) + ⟨add_mem _ hx' hy', H1 _ _ _ _ hx hy⟩) (λ r x hx, exists.elim hx $ λ hx' hx, + ⟨smul_mem _ _ hx', H2 r _ _ hx⟩) end @[simp] lemma span_span_coe_preimage : span R ((coe : span R s → M) ⁻¹' s) = ⊤ := -begin - refine eq_top_iff.2 (λ x hx, span_induction' (λ x hx, _) _ _ (λ r x hx, _) x), +eq_top_iff.2 $ λ x, subtype.rec_on x $ λ x hx _, begin + refine span_induction' (λ x hx, _) _ (λ x y _ _, _) (λ r x _, _) hx, { exact subset_span hx }, - { exact submodule.zero_mem _ }, - { intros x y hx hy, - exact submodule.add_mem _ hx hy }, - { exact submodule.smul_mem _ _ hx } + { exact zero_mem _ }, + { exact add_mem _ }, + { exact smul_mem _ _ } end lemma span_nat_eq_add_submonoid_closure (s : set M) : diff --git a/src/model_theory/basic.lean b/src/model_theory/basic.lean index 6665fa6848f9d..83e44d98bc8f9 100644 --- a/src/model_theory/basic.lean +++ b/src/model_theory/basic.lean @@ -774,21 +774,15 @@ def subtype (S : L.substructure M) : S ↪[L] M := @[simp] theorem coe_subtype : ⇑S.subtype = coe := rfl -/-- An induction principle on elements of the type `substructure.closure L s`. -If `p` holds for `1` and all elements of `s`, and is preserved under multiplication, then `p` -holds for all elements of the closure of `s`. -The difference with `substructure.closure_induction` is that this acts on the subtype. --/ -@[elab_as_eliminator] lemma closure_induction' (s : set M) {p : closure L s → Prop} - (Hs : ∀ x (h : x ∈ s), p ⟨x, subset_closure h⟩) - (Hfun : ∀ {n : ℕ} (f : L.functions n), closed_under f (set_of p)) - (x : closure L s) : - p x := -subtype.rec_on x $ λ x hx, begin - refine exists.elim _ (λ (hx : x ∈ closure L s) (hc : p ⟨x, hx⟩), hc), - exact closure_induction hx (λ x hx, ⟨subset_closure hx, Hs x hx⟩) (λ n f x hx, - ⟨(closure L s).fun_mem f _ (λ i, classical.some (hx i)), - Hfun f (λ i, ⟨x i, classical.some (hx i)⟩) (λ i, classical.some_spec (hx i))⟩), +/-- A dependent version of `substructure.closure_induction`. -/ +@[elab_as_eliminator] lemma closure_induction' (s : set M) {p : Π x, x ∈ closure L s → Prop} + (Hs : ∀ x (h : x ∈ s), p x (subset_closure h)) + (Hfun : ∀ {n : ℕ} (f : L.functions n), closed_under f {x | ∃ hx, p x hx}) + {x} (hx : x ∈ closure L s) : + p x hx := +begin + refine exists.elim _ (λ (hx : x ∈ closure L s) (hc : p x hx), hc), + exact closure_induction hx (λ x hx, ⟨subset_closure hx, Hs x hx⟩) @Hfun end end substructure