From fe03e7654edb8bfed371f6f6a654aee39c480e09 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Thu, 24 Jan 2019 18:37:05 +0000 Subject: [PATCH 01/22] stuff --- src/algebra/direct_limit.lean | 290 ++++++++++++++++++++++++++++++++++ 1 file changed, 290 insertions(+) create mode 100644 src/algebra/direct_limit.lean diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean new file mode 100644 index 0000000000000..8f9ab3a333a88 --- /dev/null +++ b/src/algebra/direct_limit.lean @@ -0,0 +1,290 @@ +import algebra.ring order.lattice + +universes u v + +open lattice + +instance is_ring_hom.to_is_add_group_hom {α : Type*} {β : Type*} [ring α] [ring β] + (f : α → β) [is_ring_hom f] : is_add_group_hom f := +⟨λ _ _, is_ring_hom.map_add f⟩ + +class directed_order (α : Type u) extends has_sup α, partial_order α := +(le_sup_left : ∀ a b : α, a ≤ a ⊔ b) +(le_sup_right : ∀ a b : α, b ≤ a ⊔ b) + +theorem le_sup_left {α : Type u} [directed_order α] {x y : α} : x ≤ x ⊔ y := +directed_order.le_sup_left x y + +theorem le_sup_right {α : Type u} [directed_order α] {x y : α} : y ≤ x ⊔ y := +directed_order.le_sup_right x y + +variables {ι : out_param (Type u)} [inhabited ι] [directed_order ι] +variables {G : ι → Type v} + +class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := +(Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) + +def direct_limit.setoid (f : Π i j, i ≤ j → G i → G j) + [directed_system f] : setoid (sigma G) := +⟨λ i j, ∃ k (hik : i.1 ≤ k) (hjk : j.1 ≤ k), f i.1 k hik i.2 = f j.1 k hjk j.2, + λ i, ⟨i.1, le_refl i.1, le_refl i.1, rfl⟩, + λ i j ⟨k, hik, hjk, H⟩, ⟨k, hjk, hik, H.symm⟩, + λ i j k ⟨ij, hiij, hjij, hij⟩ ⟨jk, hjjk, hkjk, hjk⟩, + ⟨ij ⊔ jk, le_trans hiij le_sup_left, le_trans hkjk le_sup_right, + calc f i.1 (ij ⊔ jk) _ i.2 + = f ij (ij ⊔ jk) _ (f i.1 ij _ i.2) : eq.symm $ directed_system.Hcomp f _ _ _ _ _ _ + ... = f ij (ij ⊔ jk) _ (f j.1 ij _ j.2) : congr_arg _ hij + ... = f j.1 (ij ⊔ jk) _ j.2 : directed_system.Hcomp f _ _ _ _ _ _ + ... = f jk (ij ⊔ jk) _ (f j.1 jk _ j.2) : eq.symm $ directed_system.Hcomp f _ _ _ _ _ _ + ... = f jk (ij ⊔ jk) _ (f k.1 jk _ k.2) : congr_arg _ hjk + ... = f k.1 (ij ⊔ jk) _ k.2 : directed_system.Hcomp f _ _ _ _ _ _⟩⟩ + +local attribute [instance] direct_limit.setoid + +def direct_limit (f : Π i j, i ≤ j → G i → G j) + [directed_system f] : Type (max u v) := +quotient (direct_limit.setoid f) + +namespace direct_limit + +variables (f : Π i j, i ≤ j → G i → G j) [directed_system f] + +def of (i x) : direct_limit f := +⟦⟨i, x⟩⟧ + +theorem of_f {i j hij x} : (of f j (f i j hij x)) = of f i x := +quotient.sound ⟨j, le_refl j, hij, directed_system.Hcomp f _ _ _ _ _ _⟩ + +variables {P : Type*} (g : Π i, G i → P) +variable (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) + +def rec (x : direct_limit f) : P := +quotient.lift_on x (λ i, g i.1 i.2) $ +λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨t, hpt, hqt, hpqt⟩, +calc g p xp + = g t (f p t hpt xp) : eq.symm $ Hg p t hpt xp +... = g t (f q t hqt xq) : congr_arg _ hpqt +... = g q xq : Hg q t hqt xq + +lemma rec_of {i} (x) : rec f g Hg (of f i x) = g i x := rfl + +theorem rec_unique (F : direct_limit f → P) + (HF : ∀ i x, F (of f i x) = g i x) (x) : + F x = rec f g Hg x := +quotient.induction_on x $ λ ⟨i, x⟩, HF i x + +end direct_limit + +namespace directed_system + +variables [∀ i, add_comm_group (G i)] (f : Π i j, i ≤ j → G i → G j) +variables [∀ i j h, is_add_group_hom (f i j h)] [directed_system f] + +theorem map_add {i j k m xi xj hik hjk hkm} : + f k m hkm (f i k hik xi + f j k hjk xj) = f i m (le_trans hik hkm) xi + f j m (le_trans hjk hkm) xj := +(is_add_group_hom.add _ _ _).trans $ +congr_arg₂ (+) (directed_system.Hcomp f _ _ _ _ _ _) (directed_system.Hcomp f _ _ _ _ _ _) + +theorem add {i j hij x y} : f i j hij (x + y) = f i j hij x + f i j hij y := +is_add_group_hom.add _ _ _ + +theorem zero {i j hij} : f i j hij 0 = 0 := +is_add_group_hom.zero _ + +theorem neg {i j hij x} : f i j hij (-x) = -f i j hij x := +is_add_group_hom.neg _ _ + +end directed_system + +namespace directed_system + +variables [∀ i, ring (G i)] (f : Π i j, i ≤ j → G i → G j) +variables [∀ i j h, is_ring_hom (f i j h)] [directed_system f] + +theorem map_mul {i j k m xi xj hik hjk hkm} : + f k m hkm (f i k hik xi * f j k hjk xj) = f i m (le_trans hik hkm) xi * f j m (le_trans hjk hkm) xj := +(is_ring_hom.map_mul _).trans $ +congr_arg₂ (*) (directed_system.Hcomp f _ _ _ _ _ _) (directed_system.Hcomp f _ _ _ _ _ _) + +theorem mul {i j hij x y} : f i j hij (x * y) = f i j hij x * f i j hij y := +is_ring_hom.map_mul _ + +theorem one {i j hij} : f i j hij 1 = 1 := +is_ring_hom.map_one _ + +end directed_system + +namespace direct_limit + +variables [∀ i, add_comm_group (G i)] (f : Π i j, i ≤ j → G i → G j) +variables [∀ i j h, is_add_group_hom (f i j h)] [directed_system f] + +local attribute [instance] direct_limit.setoid + +instance add_comm_group' : add_comm_group (quotient (direct_limit.setoid f)) := +{ add := λ i j, quotient.lift_on₂ i j + (λ xi xj, ⟦⟨xi.1 ⊔ xj.1, f xi.1 (xi.1 ⊔ xj.1) le_sup_left xi.2 + + f xj.1 (xi.1 ⊔ xj.1) le_sup_right xj.2⟩⟧ : sigma G → sigma G → quotient (direct_limit.setoid f)) + (λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩ ⟨s, xs⟩ ⟨t, hpt, hrt, hprt⟩ ⟨u, hqu, hsu, hqsu⟩, quotient.sound $ + ⟨(p ⊔ q) ⊔ (r ⊔ s) ⊔ (t ⊔ u), + le_trans le_sup_left le_sup_left, + le_trans le_sup_right le_sup_left, + calc f (p ⊔ q) (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans le_sup_left le_sup_left) + (f p (p ⊔ q) le_sup_left xp + f q (p ⊔ q) le_sup_right xq) + = f p (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans hpt (le_trans le_sup_left le_sup_right)) xp + + f q (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans hqu (le_trans le_sup_right le_sup_right)) xq : + directed_system.map_add f + ... = f t (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f p t hpt xp) + + f u (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f q u hqu xq) : + congr_arg₂ (+) (directed_system.Hcomp f _ _ _ _ _ _).symm (directed_system.Hcomp f _ _ _ _ _ _).symm + ... = f t (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f r t hrt xr) + + f u (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f s u hsu xs) : + congr_arg₂ (+) (congr_arg _ hprt) (congr_arg _ hqsu) + ... = f r (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ xr + + f s (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ xs : + congr_arg₂ (+) (directed_system.Hcomp f _ _ _ _ _ _) (directed_system.Hcomp f _ _ _ _ _ _) + ... = f (r ⊔ s) (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans le_sup_right le_sup_left) + (f r (r ⊔ s) le_sup_left xr + f s (r ⊔ s) le_sup_right xs) : + eq.symm $ directed_system.map_add f⟩), + add_assoc := λ i j k, quotient.induction_on₃ i j k $ λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩, quotient.sound + ⟨((p ⊔ q) ⊔ r) ⊔ (p ⊔ (q ⊔ r)), le_sup_left, le_sup_right, + by dsimp; simp [directed_system.map_add f, directed_system.add f, directed_system.Hcomp f, -add_comm]⟩, + zero := ⟦⟨default _, 0⟩⟧, + zero_add := λ i, quotient.induction_on i $ λ ⟨p, xp⟩, quotient.sound + ⟨default _ ⊔ p, le_refl _, le_sup_right, + by dsimp; simp [directed_system.map_add f, directed_system.zero f, directed_system.Hcomp f]⟩, + add_zero := λ i, quotient.induction_on i $ λ ⟨p, xp⟩, quotient.sound + ⟨p ⊔ default ι, le_refl _, le_sup_left, + by dsimp; simp [directed_system.map_add f, directed_system.zero f, directed_system.Hcomp f]⟩, + neg := λ i, quotient.lift_on i (λ p, ⟦⟨p.1, -p.2⟩⟧ : sigma G → quotient (direct_limit.setoid f)) $ + λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨t, hpt, hqt, hpqt⟩, quotient.sound $ + ⟨t, hpt, hqt, by dsimp at hpqt ⊢; simp [directed_system.neg f, hpqt]⟩, + add_left_neg := λ i, quotient.induction_on i $ λ ⟨p, xp⟩, quotient.sound $ + ⟨(p ⊔ p) ⊔ default ι, le_sup_left, le_sup_right, + by dsimp; simp [directed_system.map_add f, directed_system.neg f, directed_system.zero f]⟩, + add_comm := λ i j, quotient.induction_on₂ i j $ λ ⟨p, xp⟩ ⟨q, xq⟩, quotient.sound + ⟨(p ⊔ q) ⊔ (q ⊔ p), le_sup_left, le_sup_right, + by dsimp; simp [directed_system.map_add f]⟩ } + +instance : add_comm_group (direct_limit f) := +direct_limit.add_comm_group' f + +instance of.is_add_group_hom {i} : is_add_group_hom (direct_limit.of f i) := +is_add_group_hom.mk _ $ λ x y, quotient.sound +⟨i ⊔ i, le_sup_left, le_refl _, + by dsimp; simp [directed_system.map_add f]; simp [directed_system.add f]⟩ + +theorem of.zero_exact {i x} (H : of f i x = 0) : + ∃ p hp1, f i p hp1 x = (0 : G p) := +let ⟨p, hp1, hp2, hp3⟩ := quotient.exact H in +⟨p, hp1, hp3.trans $ is_add_group_hom.zero _⟩ + +variables {P : Type*} [add_comm_group P] +variables (g : Π i, G i → P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) +variables [∀ i, is_add_group_hom (g i)] + +def rec.is_add_group_hom : is_add_group_hom (rec f g Hg) := +is_add_group_hom.mk _ $ λ i j, quotient.induction_on₂ i j $ λ ⟨p, xp⟩ ⟨q, xq⟩, +calc _ = _ : is_add_group_hom.add _ _ _ + ... = _ : congr_arg₂ (+) (Hg _ _ _ _) (Hg _ _ _ _) + +end direct_limit + +namespace direct_limit + +variables [∀ i, ring (G i)] (f : Π i j, i ≤ j → G i → G j) +variables [∀ i j h, is_ring_hom (f i j h)] [directed_system f] + +local attribute [instance] direct_limit.setoid + +instance ring' : ring (quotient (direct_limit.setoid f)) := +{ mul := λ i j, quotient.lift_on₂ i j + (λ xi xj, ⟦⟨xi.1 ⊔ xj.1, f xi.1 (xi.1 ⊔ xj.1) le_sup_left xi.2 * + f xj.1 (xi.1 ⊔ xj.1) le_sup_right xj.2⟩⟧ : sigma G → sigma G → quotient (direct_limit.setoid f)) + (λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩ ⟨s, xs⟩ ⟨t, hpt, hrt, hprt⟩ ⟨u, hqu, hsu, hqsu⟩, quotient.sound $ + ⟨(p ⊔ q) ⊔ (r ⊔ s) ⊔ (t ⊔ u), + le_trans le_sup_left le_sup_left, + le_trans le_sup_right le_sup_left, + calc f (p ⊔ q) (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans le_sup_left le_sup_left) + (f p (p ⊔ q) le_sup_left xp * f q (p ⊔ q) le_sup_right xq) + = f p (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans hpt (le_trans le_sup_left le_sup_right)) xp + * f q (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans hqu (le_trans le_sup_right le_sup_right)) xq : + directed_system.map_mul f + ... = f t (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f p t hpt xp) + * f u (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f q u hqu xq) : + congr_arg₂ (*) (directed_system.Hcomp f _ _ _ _ _ _).symm (directed_system.Hcomp f _ _ _ _ _ _).symm + ... = f t (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f r t hrt xr) + * f u (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f s u hsu xs) : + congr_arg₂ (*) (congr_arg _ hprt) (congr_arg _ hqsu) + ... = f r (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ xr + * f s (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ xs : + congr_arg₂ (*) (directed_system.Hcomp f _ _ _ _ _ _) (directed_system.Hcomp f _ _ _ _ _ _) + ... = f (r ⊔ s) (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans le_sup_right le_sup_left) + (f r (r ⊔ s) le_sup_left xr * f s (r ⊔ s) le_sup_right xs) : + eq.symm $ directed_system.map_mul f⟩), + mul_assoc := λ i j k, quotient.induction_on₃ i j k $ λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩, quotient.sound + ⟨((p ⊔ q) ⊔ r) ⊔ (p ⊔ (q ⊔ r)), le_sup_left, le_sup_right, + by dsimp; simp [directed_system.map_mul f, directed_system.mul f, directed_system.Hcomp f, mul_assoc]⟩, + one := ⟦⟨default _, 1⟩⟧, + one_mul := λ i, quotient.induction_on i $ λ ⟨p, xp⟩, quotient.sound + ⟨default _ ⊔ p, le_refl _, le_sup_right, + by dsimp; simp [directed_system.map_mul f, directed_system.one f, directed_system.Hcomp f]⟩, + mul_one := λ i, quotient.induction_on i $ λ ⟨p, xp⟩, quotient.sound + ⟨p ⊔ default ι, le_refl _, le_sup_left, + by dsimp; simp [directed_system.map_mul f, directed_system.one f, directed_system.Hcomp f]⟩, + left_distrib := λ i j k, quotient.induction_on₃ i j k $ λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩, quotient.sound + ⟨p ⊔ (q ⊔ r) ⊔ (p ⊔ q ⊔ (p ⊔ r)), le_sup_left, le_sup_right, + by dsimp; simp [directed_system.add f, directed_system.mul f, directed_system.Hcomp f, mul_add]⟩, + right_distrib := λ i j k, quotient.induction_on₃ i j k $ λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩, quotient.sound + ⟨p ⊔ q ⊔ r ⊔ (p ⊔ r ⊔ (q ⊔ r)), le_sup_left, le_sup_right, + by dsimp; simp [directed_system.add f, directed_system.mul f, directed_system.Hcomp f, add_mul]⟩, + .. direct_limit.add_comm_group' f } + +instance: ring (direct_limit f) := +direct_limit.ring' f + +instance of.is_ring_hom {i} : is_ring_hom (direct_limit.of f i) := +{ map_add := is_add_group_hom.add _, + map_mul := λ x y, quotient.sound ⟨i ⊔ i, le_sup_left, le_refl _, + by dsimp; simp [directed_system.mul f, directed_system.Hcomp f]⟩, + map_one := quotient.sound ⟨i ⊔ default _, le_sup_left, le_sup_right, + by dsimp; simp [directed_system.one f]⟩ } + +theorem of.one_exact {i x} (H : of f i x = 1) : + ∃ p hp1, f i p hp1 x = (1 : G p) := +let ⟨p, hp1, hp2, hp3⟩ := quotient.exact H in +⟨p, hp1, hp3.trans $ is_ring_hom.map_one _⟩ + +variables {P : Type*} [ring P] +variables (g : Π i, G i → P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) +variables [∀ i, is_ring_hom (g i)] + +local attribute [instance] rec.is_add_group_hom + +def rec.is_ring_hom : is_ring_hom (rec f g Hg) := +{ map_add := is_add_group_hom.add _, + map_mul := λ i j, quotient.induction_on₂ i j $ λ ⟨p, xp⟩ ⟨q, xq⟩, + calc _ = _ : is_ring_hom.map_mul _ + ... = _ : congr_arg₂ (*) (Hg _ _ _ _) (Hg _ _ _ _), + map_one := show g (default _) 1 = 1, from is_ring_hom.map_one _ } + +end direct_limit + +namespace direct_limit + +variables [∀ i, comm_ring (G i)] (f : Π i j, i ≤ j → G i → G j) +variables [∀ i j h, is_ring_hom (f i j h)] [directed_system f] + +local attribute [instance] direct_limit.setoid + +instance comm_ring' : comm_ring (quotient (direct_limit.setoid f)) := +{ + mul_comm := λ i j, quotient.induction_on₂ i j $ λ ⟨p, xp⟩ ⟨q, xq⟩, quotient.sound + ⟨(p ⊔ q) ⊔ (q ⊔ p), le_sup_left, le_sup_right, + by dsimp; simp [directed_system.map_mul f, directed_system.mul f, directed_system.Hcomp f, mul_comm]⟩ + .. direct_limit.ring' f } + +instance: comm_ring (direct_limit f) := +direct_limit.comm_ring' f + +end direct_limit From 4b68946f393e90c49eb7e3bce5a9abd7af30f723 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Thu, 24 Jan 2019 18:29:12 +0000 Subject: [PATCH 02/22] stuff --- src/algebra/direct_limit.lean | 41 ++++++++++++++++++++++++----------- 1 file changed, 28 insertions(+), 13 deletions(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 8f9ab3a333a88..4c2f98b329408 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -1,18 +1,35 @@ -import algebra.ring order.lattice - -universes u v +import linear_algebra.direct_sum_module +universes u v w +#check directed +#check @directed.finset_le +#check submodule.Union_coe_of_directed +#check dfinsupp open lattice -instance is_ring_hom.to_is_add_group_hom {α : Type*} {β : Type*} [ring α] [ring β] - (f : α → β) [is_ring_hom f] : is_add_group_hom f := -⟨λ _ _, is_ring_hom.map_add f⟩ +class directed_order (α : Type u) extends partial_order α := +(directed : ∀ i j : α, ∃ k, i ≤ k ∧ j ≤ k) + +variables {ι : out_param (Type u)} [inhabited ι] +variables [directed_order ι] [decidable_eq ι] +variables {G : ι → Type v} + +class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := +(Hid : ∀ i x, f i i (le_refl i) x = x) +(Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) -class directed_order (α : Type u) extends has_sup α, partial_order α := -(le_sup_left : ∀ a b : α, a ≤ a ⊔ b) -(le_sup_right : ∀ a b : α, b ≤ a ⊔ b) +section module +variables [Π i, decidable_eq (G i)] {R : Type w} [comm_ring R] +variables [Π i, add_comm_group (G i)] [Π i, module R (G i)] -theorem le_sup_left {α : Type u} [directed_order α] {x y : α} : x ≤ x ⊔ y := +def direct_limit.submodule (f : Π i j, i ≤ j → G i →ₗ G j) : + submodule R (direct_sum R ι G) := +submodule.span $ ⋃ i j (H : i ≤ j) x, +{ @direct_sum.of _ _ ι _ G _ _ i x - direct_sum.of G j (f i j H x) } + +end module + +/-theorem le_sup_left {α : Type u} [directed_order α] {x y : α} : x ≤ x ⊔ y := directed_order.le_sup_left x y theorem le_sup_right {α : Type u} [directed_order α] {x y : α} : y ≤ x ⊔ y := @@ -21,9 +38,6 @@ directed_order.le_sup_right x y variables {ι : out_param (Type u)} [inhabited ι] [directed_order ι] variables {G : ι → Type v} -class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := -(Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) - def direct_limit.setoid (f : Π i j, i ≤ j → G i → G j) [directed_system f] : setoid (sigma G) := ⟨λ i j, ∃ k (hik : i.1 ≤ k) (hjk : j.1 ≤ k), f i.1 k hik i.2 = f j.1 k hjk j.2, @@ -288,3 +302,4 @@ instance: comm_ring (direct_limit f) := direct_limit.comm_ring' f end direct_limit +-/ \ No newline at end of file From 696a98456fd084f522252a09509dbd3558163fb9 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Thu, 24 Jan 2019 22:09:42 +0000 Subject: [PATCH 03/22] more stuff --- src/algebra/direct_limit.lean | 366 ++++++++-------------------------- 1 file changed, 78 insertions(+), 288 deletions(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 4c2f98b329408..2eb6be2e04b64 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -1,305 +1,95 @@ -import linear_algebra.direct_sum_module +import linear_algebra.direct_sum_module linear_algebra.linear_combination -universes u v w -#check directed -#check @directed.finset_le -#check submodule.Union_coe_of_directed -#check dfinsupp -open lattice +universes u v w u₁ + +open lattice submodule class directed_order (α : Type u) extends partial_order α := (directed : ∀ i j : α, ∃ k, i ≤ k ∧ j ≤ k) -variables {ι : out_param (Type u)} [inhabited ι] +variables {ι : Type u} [inhabited ι] variables [directed_order ι] [decidable_eq ι] -variables {G : ι → Type v} +variables (G : ι → Type v) [Π i, decidable_eq (G i)] -class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := +class directed_system {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] + (f : Π i j, i ≤ j → G i →ₗ G j) : Prop := (Hid : ∀ i x, f i i (le_refl i) x = x) (Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) -section module -variables [Π i, decidable_eq (G i)] {R : Type w} [comm_ring R] -variables [Π i, add_comm_group (G i)] [Π i, module R (G i)] - -def direct_limit.submodule (f : Π i j, i ≤ j → G i →ₗ G j) : - submodule R (direct_sum R ι G) := -submodule.span $ ⋃ i j (H : i ≤ j) x, -{ @direct_sum.of _ _ ι _ G _ _ i x - direct_sum.of G j (f i j H x) } - -end module - -/-theorem le_sup_left {α : Type u} [directed_order α] {x y : α} : x ≤ x ⊔ y := -directed_order.le_sup_left x y - -theorem le_sup_right {α : Type u} [directed_order α] {x y : α} : y ≤ x ⊔ y := -directed_order.le_sup_right x y - -variables {ι : out_param (Type u)} [inhabited ι] [directed_order ι] -variables {G : ι → Type v} - -def direct_limit.setoid (f : Π i j, i ≤ j → G i → G j) - [directed_system f] : setoid (sigma G) := -⟨λ i j, ∃ k (hik : i.1 ≤ k) (hjk : j.1 ≤ k), f i.1 k hik i.2 = f j.1 k hjk j.2, - λ i, ⟨i.1, le_refl i.1, le_refl i.1, rfl⟩, - λ i j ⟨k, hik, hjk, H⟩, ⟨k, hjk, hik, H.symm⟩, - λ i j k ⟨ij, hiij, hjij, hij⟩ ⟨jk, hjjk, hkjk, hjk⟩, - ⟨ij ⊔ jk, le_trans hiij le_sup_left, le_trans hkjk le_sup_right, - calc f i.1 (ij ⊔ jk) _ i.2 - = f ij (ij ⊔ jk) _ (f i.1 ij _ i.2) : eq.symm $ directed_system.Hcomp f _ _ _ _ _ _ - ... = f ij (ij ⊔ jk) _ (f j.1 ij _ j.2) : congr_arg _ hij - ... = f j.1 (ij ⊔ jk) _ j.2 : directed_system.Hcomp f _ _ _ _ _ _ - ... = f jk (ij ⊔ jk) _ (f j.1 jk _ j.2) : eq.symm $ directed_system.Hcomp f _ _ _ _ _ _ - ... = f jk (ij ⊔ jk) _ (f k.1 jk _ k.2) : congr_arg _ hjk - ... = f k.1 (ij ⊔ jk) _ k.2 : directed_system.Hcomp f _ _ _ _ _ _⟩⟩ - -local attribute [instance] direct_limit.setoid - -def direct_limit (f : Π i j, i ≤ j → G i → G j) - [directed_system f] : Type (max u v) := -quotient (direct_limit.setoid f) +def direct_limit {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] + (f : Π i j, i ≤ j → G i →ₗ G j) [directed_system G f] : Type (max u v) := +quotient $ span $ ⋃ i j (H : i ≤ j) x, +({ direct_sum.of ι G i x - direct_sum.of ι G j (f i j H x) } : set $ direct_sum R ι G) namespace direct_limit -variables (f : Π i j, i ≤ j → G i → G j) [directed_system f] - -def of (i x) : direct_limit f := -⟦⟨i, x⟩⟧ - -theorem of_f {i j hij x} : (of f j (f i j hij x)) = of f i x := -quotient.sound ⟨j, le_refl j, hij, directed_system.Hcomp f _ _ _ _ _ _⟩ - -variables {P : Type*} (g : Π i, G i → P) -variable (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) - -def rec (x : direct_limit f) : P := -quotient.lift_on x (λ i, g i.1 i.2) $ -λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨t, hpt, hqt, hpqt⟩, -calc g p xp - = g t (f p t hpt xp) : eq.symm $ Hg p t hpt xp -... = g t (f q t hqt xq) : congr_arg _ hpqt -... = g q xq : Hg q t hqt xq - -lemma rec_of {i} (x) : rec f g Hg (of f i x) = g i x := rfl - -theorem rec_unique (F : direct_limit f → P) - (HF : ∀ i x, F (of f i x) = g i x) (x) : - F x = rec f g Hg x := -quotient.induction_on x $ λ ⟨i, x⟩, HF i x - -end direct_limit - -namespace directed_system - -variables [∀ i, add_comm_group (G i)] (f : Π i j, i ≤ j → G i → G j) -variables [∀ i j h, is_add_group_hom (f i j h)] [directed_system f] - -theorem map_add {i j k m xi xj hik hjk hkm} : - f k m hkm (f i k hik xi + f j k hjk xj) = f i m (le_trans hik hkm) xi + f j m (le_trans hjk hkm) xj := -(is_add_group_hom.add _ _ _).trans $ -congr_arg₂ (+) (directed_system.Hcomp f _ _ _ _ _ _) (directed_system.Hcomp f _ _ _ _ _ _) - -theorem add {i j hij x y} : f i j hij (x + y) = f i j hij x + f i j hij y := -is_add_group_hom.add _ _ _ - -theorem zero {i j hij} : f i j hij 0 = 0 := -is_add_group_hom.zero _ - -theorem neg {i j hij x} : f i j hij (-x) = -f i j hij x := -is_add_group_hom.neg _ _ - -end directed_system - -namespace directed_system - -variables [∀ i, ring (G i)] (f : Π i j, i ≤ j → G i → G j) -variables [∀ i j h, is_ring_hom (f i j h)] [directed_system f] - -theorem map_mul {i j k m xi xj hik hjk hkm} : - f k m hkm (f i k hik xi * f j k hjk xj) = f i m (le_trans hik hkm) xi * f j m (le_trans hjk hkm) xj := -(is_ring_hom.map_mul _).trans $ -congr_arg₂ (*) (directed_system.Hcomp f _ _ _ _ _ _) (directed_system.Hcomp f _ _ _ _ _ _) - -theorem mul {i j hij x y} : f i j hij (x * y) = f i j hij x * f i j hij y := -is_ring_hom.map_mul _ - -theorem one {i j hij} : f i j hij 1 = 1 := -is_ring_hom.map_one _ - -end directed_system - -namespace direct_limit - -variables [∀ i, add_comm_group (G i)] (f : Π i j, i ≤ j → G i → G j) -variables [∀ i j h, is_add_group_hom (f i j h)] [directed_system f] - -local attribute [instance] direct_limit.setoid - -instance add_comm_group' : add_comm_group (quotient (direct_limit.setoid f)) := -{ add := λ i j, quotient.lift_on₂ i j - (λ xi xj, ⟦⟨xi.1 ⊔ xj.1, f xi.1 (xi.1 ⊔ xj.1) le_sup_left xi.2 + - f xj.1 (xi.1 ⊔ xj.1) le_sup_right xj.2⟩⟧ : sigma G → sigma G → quotient (direct_limit.setoid f)) - (λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩ ⟨s, xs⟩ ⟨t, hpt, hrt, hprt⟩ ⟨u, hqu, hsu, hqsu⟩, quotient.sound $ - ⟨(p ⊔ q) ⊔ (r ⊔ s) ⊔ (t ⊔ u), - le_trans le_sup_left le_sup_left, - le_trans le_sup_right le_sup_left, - calc f (p ⊔ q) (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans le_sup_left le_sup_left) - (f p (p ⊔ q) le_sup_left xp + f q (p ⊔ q) le_sup_right xq) - = f p (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans hpt (le_trans le_sup_left le_sup_right)) xp - + f q (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans hqu (le_trans le_sup_right le_sup_right)) xq : - directed_system.map_add f - ... = f t (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f p t hpt xp) - + f u (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f q u hqu xq) : - congr_arg₂ (+) (directed_system.Hcomp f _ _ _ _ _ _).symm (directed_system.Hcomp f _ _ _ _ _ _).symm - ... = f t (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f r t hrt xr) - + f u (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f s u hsu xs) : - congr_arg₂ (+) (congr_arg _ hprt) (congr_arg _ hqsu) - ... = f r (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ xr - + f s (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ xs : - congr_arg₂ (+) (directed_system.Hcomp f _ _ _ _ _ _) (directed_system.Hcomp f _ _ _ _ _ _) - ... = f (r ⊔ s) (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans le_sup_right le_sup_left) - (f r (r ⊔ s) le_sup_left xr + f s (r ⊔ s) le_sup_right xs) : - eq.symm $ directed_system.map_add f⟩), - add_assoc := λ i j k, quotient.induction_on₃ i j k $ λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩, quotient.sound - ⟨((p ⊔ q) ⊔ r) ⊔ (p ⊔ (q ⊔ r)), le_sup_left, le_sup_right, - by dsimp; simp [directed_system.map_add f, directed_system.add f, directed_system.Hcomp f, -add_comm]⟩, - zero := ⟦⟨default _, 0⟩⟧, - zero_add := λ i, quotient.induction_on i $ λ ⟨p, xp⟩, quotient.sound - ⟨default _ ⊔ p, le_refl _, le_sup_right, - by dsimp; simp [directed_system.map_add f, directed_system.zero f, directed_system.Hcomp f]⟩, - add_zero := λ i, quotient.induction_on i $ λ ⟨p, xp⟩, quotient.sound - ⟨p ⊔ default ι, le_refl _, le_sup_left, - by dsimp; simp [directed_system.map_add f, directed_system.zero f, directed_system.Hcomp f]⟩, - neg := λ i, quotient.lift_on i (λ p, ⟦⟨p.1, -p.2⟩⟧ : sigma G → quotient (direct_limit.setoid f)) $ - λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨t, hpt, hqt, hpqt⟩, quotient.sound $ - ⟨t, hpt, hqt, by dsimp at hpqt ⊢; simp [directed_system.neg f, hpqt]⟩, - add_left_neg := λ i, quotient.induction_on i $ λ ⟨p, xp⟩, quotient.sound $ - ⟨(p ⊔ p) ⊔ default ι, le_sup_left, le_sup_right, - by dsimp; simp [directed_system.map_add f, directed_system.neg f, directed_system.zero f]⟩, - add_comm := λ i j, quotient.induction_on₂ i j $ λ ⟨p, xp⟩ ⟨q, xq⟩, quotient.sound - ⟨(p ⊔ q) ⊔ (q ⊔ p), le_sup_left, le_sup_right, - by dsimp; simp [directed_system.map_add f]⟩ } - -instance : add_comm_group (direct_limit f) := -direct_limit.add_comm_group' f - -instance of.is_add_group_hom {i} : is_add_group_hom (direct_limit.of f i) := -is_add_group_hom.mk _ $ λ x y, quotient.sound -⟨i ⊔ i, le_sup_left, le_refl _, - by dsimp; simp [directed_system.map_add f]; simp [directed_system.add f]⟩ - -theorem of.zero_exact {i x} (H : of f i x = 0) : - ∃ p hp1, f i p hp1 x = (0 : G p) := -let ⟨p, hp1, hp2, hp3⟩ := quotient.exact H in -⟨p, hp1, hp3.trans $ is_add_group_hom.zero _⟩ - -variables {P : Type*} [add_comm_group P] -variables (g : Π i, G i → P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) -variables [∀ i, is_add_group_hom (g i)] - -def rec.is_add_group_hom : is_add_group_hom (rec f g Hg) := -is_add_group_hom.mk _ $ λ i j, quotient.induction_on₂ i j $ λ ⟨p, xp⟩ ⟨q, xq⟩, -calc _ = _ : is_add_group_hom.add _ _ _ - ... = _ : congr_arg₂ (+) (Hg _ _ _ _) (Hg _ _ _ _) - -end direct_limit - -namespace direct_limit - -variables [∀ i, ring (G i)] (f : Π i j, i ≤ j → G i → G j) -variables [∀ i j h, is_ring_hom (f i j h)] [directed_system f] - -local attribute [instance] direct_limit.setoid - -instance ring' : ring (quotient (direct_limit.setoid f)) := -{ mul := λ i j, quotient.lift_on₂ i j - (λ xi xj, ⟦⟨xi.1 ⊔ xj.1, f xi.1 (xi.1 ⊔ xj.1) le_sup_left xi.2 * - f xj.1 (xi.1 ⊔ xj.1) le_sup_right xj.2⟩⟧ : sigma G → sigma G → quotient (direct_limit.setoid f)) - (λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩ ⟨s, xs⟩ ⟨t, hpt, hrt, hprt⟩ ⟨u, hqu, hsu, hqsu⟩, quotient.sound $ - ⟨(p ⊔ q) ⊔ (r ⊔ s) ⊔ (t ⊔ u), - le_trans le_sup_left le_sup_left, - le_trans le_sup_right le_sup_left, - calc f (p ⊔ q) (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans le_sup_left le_sup_left) - (f p (p ⊔ q) le_sup_left xp * f q (p ⊔ q) le_sup_right xq) - = f p (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans hpt (le_trans le_sup_left le_sup_right)) xp - * f q (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans hqu (le_trans le_sup_right le_sup_right)) xq : - directed_system.map_mul f - ... = f t (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f p t hpt xp) - * f u (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f q u hqu xq) : - congr_arg₂ (*) (directed_system.Hcomp f _ _ _ _ _ _).symm (directed_system.Hcomp f _ _ _ _ _ _).symm - ... = f t (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f r t hrt xr) - * f u (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ (f s u hsu xs) : - congr_arg₂ (*) (congr_arg _ hprt) (congr_arg _ hqsu) - ... = f r (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ xr - * f s (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) _ xs : - congr_arg₂ (*) (directed_system.Hcomp f _ _ _ _ _ _) (directed_system.Hcomp f _ _ _ _ _ _) - ... = f (r ⊔ s) (p ⊔ q ⊔ (r ⊔ s) ⊔ (t ⊔ u)) (le_trans le_sup_right le_sup_left) - (f r (r ⊔ s) le_sup_left xr * f s (r ⊔ s) le_sup_right xs) : - eq.symm $ directed_system.map_mul f⟩), - mul_assoc := λ i j k, quotient.induction_on₃ i j k $ λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩, quotient.sound - ⟨((p ⊔ q) ⊔ r) ⊔ (p ⊔ (q ⊔ r)), le_sup_left, le_sup_right, - by dsimp; simp [directed_system.map_mul f, directed_system.mul f, directed_system.Hcomp f, mul_assoc]⟩, - one := ⟦⟨default _, 1⟩⟧, - one_mul := λ i, quotient.induction_on i $ λ ⟨p, xp⟩, quotient.sound - ⟨default _ ⊔ p, le_refl _, le_sup_right, - by dsimp; simp [directed_system.map_mul f, directed_system.one f, directed_system.Hcomp f]⟩, - mul_one := λ i, quotient.induction_on i $ λ ⟨p, xp⟩, quotient.sound - ⟨p ⊔ default ι, le_refl _, le_sup_left, - by dsimp; simp [directed_system.map_mul f, directed_system.one f, directed_system.Hcomp f]⟩, - left_distrib := λ i j k, quotient.induction_on₃ i j k $ λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩, quotient.sound - ⟨p ⊔ (q ⊔ r) ⊔ (p ⊔ q ⊔ (p ⊔ r)), le_sup_left, le_sup_right, - by dsimp; simp [directed_system.add f, directed_system.mul f, directed_system.Hcomp f, mul_add]⟩, - right_distrib := λ i j k, quotient.induction_on₃ i j k $ λ ⟨p, xp⟩ ⟨q, xq⟩ ⟨r, xr⟩, quotient.sound - ⟨p ⊔ q ⊔ r ⊔ (p ⊔ r ⊔ (q ⊔ r)), le_sup_left, le_sup_right, - by dsimp; simp [directed_system.add f, directed_system.mul f, directed_system.Hcomp f, add_mul]⟩, - .. direct_limit.add_comm_group' f } - -instance: ring (direct_limit f) := -direct_limit.ring' f - -instance of.is_ring_hom {i} : is_ring_hom (direct_limit.of f i) := -{ map_add := is_add_group_hom.add _, - map_mul := λ x y, quotient.sound ⟨i ⊔ i, le_sup_left, le_refl _, - by dsimp; simp [directed_system.mul f, directed_system.Hcomp f]⟩, - map_one := quotient.sound ⟨i ⊔ default _, le_sup_left, le_sup_right, - by dsimp; simp [directed_system.one f]⟩ } - -theorem of.one_exact {i x} (H : of f i x = 1) : - ∃ p hp1, f i p hp1 x = (1 : G p) := -let ⟨p, hp1, hp2, hp3⟩ := quotient.exact H in -⟨p, hp1, hp3.trans $ is_ring_hom.map_one _⟩ - -variables {P : Type*} [ring P] -variables (g : Π i, G i → P) (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) -variables [∀ i, is_ring_hom (g i)] - -local attribute [instance] rec.is_add_group_hom - -def rec.is_ring_hom : is_ring_hom (rec f g Hg) := -{ map_add := is_add_group_hom.add _, - map_mul := λ i j, quotient.induction_on₂ i j $ λ ⟨p, xp⟩ ⟨q, xq⟩, - calc _ = _ : is_ring_hom.map_mul _ - ... = _ : congr_arg₂ (*) (Hg _ _ _ _) (Hg _ _ _ _), - map_one := show g (default _) 1 = 1, from is_ring_hom.map_one _ } - -end direct_limit - -namespace direct_limit - -variables [∀ i, comm_ring (G i)] (f : Π i j, i ≤ j → G i → G j) -variables [∀ i j h, is_ring_hom (f i j h)] [directed_system f] - -local attribute [instance] direct_limit.setoid +section module -instance comm_ring' : comm_ring (quotient (direct_limit.setoid f)) := -{ - mul_comm := λ i j, quotient.induction_on₂ i j $ λ ⟨p, xp⟩ ⟨q, xq⟩, quotient.sound - ⟨(p ⊔ q) ⊔ (q ⊔ p), le_sup_left, le_sup_right, - by dsimp; simp [directed_system.map_mul f, directed_system.mul f, directed_system.Hcomp f, mul_comm]⟩ - .. direct_limit.ring' f } +variables {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] +variables (f : Π i j, i ≤ j → G i →ₗ G j) [directed_system G f] +include R + +instance : add_comm_group (direct_limit G f) := quotient.add_comm_group _ +instance : module R (direct_limit G f) := quotient.module _ + +variables (ι R) +def of (i) : G i →ₗ direct_limit G f := +(mkq _).comp $ direct_sum.of ι G i +variables {ι R} + +theorem of_f {i j hij x} : (of ι G R f j (f i j hij x)) = of ι G R f i x := +eq.symm $ (submodule.quotient.eq _).2 $ subset_span $ +set.mem_Union.2 ⟨i, set.mem_Union.2 ⟨j, set.mem_Union.2 ⟨hij, set.mem_Union.2 ⟨x, or.inl rfl⟩⟩⟩⟩ + +variables {P : Type u₁} [add_comm_group P] [module R P] (g : Π i, G i →ₗ P) +variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) +include Hg + +variables (ι R) +def rec : direct_limit G f →ₗ P := +liftq _ (direct_sum.to_module _ g) $ span_le.2 $ set.Union_subset $ λ i, +set.Union_subset $ λ j, set.Union_subset $ λ hij, set.Union_subset $ λ x, +set.singleton_subset_iff.2 $ linear_map.sub_mem_ker_iff.2 $ +by rw [direct_sum.to_module.of, direct_sum.to_module.of, Hg] +variables {ι R} + +omit Hg +lemma rec_of {i} (x) : rec ι G R f g Hg (of ι G R f i x) = g i x := +direct_sum.to_module.of _ _ + +theorem rec_unique (F : direct_limit G f →ₗ P) (x) : + F x = rec ι G R f (λ i, F.comp $ of ι G R f i) + (λ i j hij x, by rw [linear_map.comp_apply, of_f]; refl) x := +quotient.induction_on' x $ λ y, direct_sum.induction_on y + ((linear_map.map_zero _).trans (linear_map.map_zero _).symm) + (λ i x, eq.symm $ rec_of _ _ _ _ _) + (λ x y ihx ihy, (linear_map.map_add F (quotient.mk' x) (quotient.mk' y)).trans $ + ihx.symm ▸ ihy.symm ▸ (linear_map.map_add _ _ _).symm) + +theorem exists_of (z : direct_limit G f) : ∃ i x, z = of ι G R f i x := +quotient.induction_on' z $ λ z, direct_sum.induction_on z + ⟨default _, 0, (linear_map.map_zero _).symm⟩ + (λ i x, ⟨i, x, rfl⟩) + (λ p q ⟨i, x, ihx⟩ ⟨j, y, ihy⟩, let ⟨k, hik, hjk⟩ := directed_order.directed i j in + ⟨k, f i k hik x + f j k hjk y, by rw [linear_map.map_add, of_f, of_f, ← ihx, ← ihy]; refl⟩) + +@[elab_as_eliminator] theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit G f) + (ih : ∀ i x, C (of ι G R f i x)) : C z := +let ⟨i, x, h⟩ := exists_of G f z in h.symm ▸ ih i x + +theorem of.zero_exact {i x} (H : of ι G R f i x = 0) : + ∃ j hij, f i j hij x = (0 : G j) := +begin + rcases mem_span_iff_lc.1 ((submodule.quotient.eq _).1 H) with ⟨l, h1, h2⟩, rw sub_zero at h2, clear H, + revert l i x, intro l, haveI : decidable_eq R := classical.dec_eq R, + apply finsupp.induction l, + { intros i x h1 h2, rw [linear_map.map_zero, ← linear_map.map_zero (direct_sum.of ι G i)] at h2, + have := direct_sum.of_inj _ h2, subst this, + exact ⟨i, le_refl i, linear_map.map_zero _⟩ }, + intros a b f haf hb0 ih i x h1 h2, + sorry +end -instance: comm_ring (direct_limit f) := -direct_limit.comm_ring' f +end module end direct_limit --/ \ No newline at end of file From 36675442ea4d766e5f983f218b44ab1eff631496 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Wed, 30 Jan 2019 12:38:01 +0000 Subject: [PATCH 04/22] pre merge commit --- src/algebra/direct_limit.lean | 214 ++++++++++++++++++++-- src/data/dfinsupp.lean | 7 + src/linear_algebra/direct_sum_module.lean | 30 +++ 3 files changed, 234 insertions(+), 17 deletions(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 2eb6be2e04b64..43adc8ede9123 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -1,7 +1,18 @@ -import linear_algebra.direct_sum_module linear_algebra.linear_combination +import linear_algebra.direct_sum_module linear_algebra.linear_combination tactic.linarith universes u v w u₁ +namespace tactic.interactive +open tactic +meta def clear_aux_decl : tactic unit := +do l ← local_context, +match l with +| [] := skip +| (e::l) := cond e.is_aux_decl (tactic.clear e) skip +end + +end tactic.interactive + open lattice submodule class directed_order (α : Type u) extends partial_order α := @@ -16,10 +27,21 @@ class directed_system {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [ (Hid : ∀ i x, f i i (le_refl i) x = x) (Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) +def thing {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] + (f : Π i j, i ≤ j → G i →ₗ G j) [directed_system G f] : set (direct_sum R ι G) := + ⋃ i j (H : i ≤ j) x, + ({ direct_sum.of ι G i x - direct_sum.of ι G j (f i j H x) } : set $ direct_sum R ι G) + def direct_limit {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] (f : Π i j, i ≤ j → G i →ₗ G j) [directed_system G f] : Type (max u v) := -quotient $ span $ ⋃ i j (H : i ≤ j) x, -({ direct_sum.of ι G i x - direct_sum.of ι G j (f i j H x) } : set $ direct_sum R ι G) +quotient $ span $ thing G f + +lemma mem_thing {G : ι → Type v} [Π i, decidable_eq (G i)] + {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] + {f : Π i j, i ≤ j → G i →ₗ G j} [directed_system G f] {a : direct_sum R ι G} : + a ∈ thing G f ↔ ∃ (i j) (H : i ≤ j) x, + a = direct_sum.of ι G i x - direct_sum.of ι G j (f i j H x) := + by simp [thing, set.mem_Union] namespace direct_limit @@ -47,10 +69,12 @@ include Hg variables (ι R) def rec : direct_limit G f →ₗ P := -liftq _ (direct_sum.to_module _ g) $ span_le.2 $ set.Union_subset $ λ i, -set.Union_subset $ λ j, set.Union_subset $ λ hij, set.Union_subset $ λ x, -set.singleton_subset_iff.2 $ linear_map.sub_mem_ker_iff.2 $ -by rw [direct_sum.to_module.of, direct_sum.to_module.of, Hg] +liftq _ (direct_sum.to_module _ g) + (span_le.2 $ set.Union_subset $ λ i, + set.Union_subset $ λ j, set.Union_subset $ λ hij, set.Union_subset $ λ x, + set.singleton_subset_iff.2 $ linear_map.sub_mem_ker_iff.2 $ + by rw [direct_sum.to_module.of, direct_sum.to_module.of, Hg]) + variables {ι R} omit Hg @@ -77,18 +101,174 @@ quotient.induction_on' z $ λ z, direct_sum.induction_on z (ih : ∀ i x, C (of ι G R f i x)) : C z := let ⟨i, x, h⟩ := exists_of G f z in h.symm ▸ ih i x +lemma span_preimage_le_comap_span {R M N: Type*} [ring R] [add_comm_group M] [module R M] + [add_comm_group N] [module R N] (f : N →ₗ M) (s : set M) : span (f ⁻¹' s) ≤ (span s).comap f := +λ x h, span_induction h + (by simp only [set.preimage, set.mem_set_of_eq, mem_comap]; exact λ x h, subset_span h) + (zero_mem ((span s).comap f)) + (λ _ _ hx hy, add_mem ((span s).comap f) hx hy) + (λ _ _ h, smul_mem ((span s).comap f) _ h) + +local attribute [instance, priority 0] classical.dec + +noncomputable def totalize : Π i j, G i →ₗ G j := +λ i j, if h : i ≤ j then f i j h else 0 + +lemma totalize_apply (i j x) : + totalize G f i j x = if h : i ≤ j then f i j h x else 0 := +if h : i ≤ j + then by dsimp [totalize]; rw [dif_pos h, dif_pos h] + else by dsimp [totalize]; rw [dif_neg h, dif_neg h, linear_map.zero_apply] +-- set_option trace.simplify.rewrite true + + + +theorem of.zero_exact {i x} (H : of ι G R f i x = 0) : + ∃ j hij, f i j hij x = (0 : G j) := +@span_induction _ _ _ _ _ _ _ + (λ y, ∀ i (x : G i), y = direct_sum.of ι G i x → + ∃ j hij, f i j hij x = (0 : G j)) + ((submodule.quotient.eq _).1 H) + (λ y hy', let ⟨i, j, Hij, z, hy⟩ := mem_thing.1 hy' in + λ k xk hyk, + let ⟨l, hl⟩ := directed_order.directed i k in + ⟨l, hl.2, begin + clear_aux_decl, + subst hyk, + have := congr_arg (direct_sum.component ι G k) hy, + clear hy', + simp at this, + subst this, + + + + end⟩) + _ + begin + simp, + + end + _ + _ _ (sub_zero _) + +theorem of.zero_exact {i x} (H : of ι G R f i x = 0) : + ∃ j hij, f i j hij x = (0 : G j) := +@span_induction _ _ _ _ _ _ _ + (λ y, ∀ i (x : G i), (direct_sum.to_module ι (λ j, totalize G f j i)) y = x → + ∃ j hij, f i j hij x = (0 : G j)) + ((submodule.quotient.eq _).1 H) + (λ y hy', let ⟨i, j, Hij, z, hy⟩ := mem_thing.1 hy' in + (λ k xk hyk, + let ⟨l, hl⟩ := directed_order.directed j k in + ⟨l, hl.2, begin + clear _let_match _let_match, + subst hyk, + have := congr_arg ((direct_sum.of ι G l).comp (direct_sum.to_module ι (λ i, totalize G f i l))) hy, + simp only [direct_sum.to_module.of, linear_map.map_sub] at this, + simp [totalize_apply] at *, + rw [dif_pos (le_trans Hij hl.1), dif_pos hl.1, directed_system.Hcomp f, + add_neg_self] at this, + clear H hy' x hy Hij z i i, + rw [← direct_sum.sum_of y, dfinsupp.sum] at *, + refine direct_sum.of_inj l _, + simp only [linear_map.map_sum, direct_sum.to_module.of, + direct_sum.ext_iff, linear_map.map_zero] at *, + assume i, + calc finset.sum (dfinsupp.support y) + (λ (m : ι), direct_sum.component ι G i + ((direct_sum.of ι G l) ((f k l _) ((totalize G f m k) (direct_sum.component ι G m y))))) + = direct_sum.component ι G i + ((direct_sum.of ι G l) ((f k l _) ((totalize G f l k) (direct_sum.component ι G l y)))) : + finset.sum_eq_single _ begin + intros, + simp [direct_sum.apply_eq_component, totalize_apply], + + end _ + ... = _ : sorry + + + + end⟩)) + (λ i x hix, ⟨i, le_refl _, + by rw [← linear_map.map_zero (direct_sum.of ι G i), direct_sum.to_module.of, + linear_map.map_zero] at hix; + rw [directed_system.Hid f, hix]⟩) + (λ x y hx hy i _ hxy, + let ⟨j, hij, hj⟩ := hx i _ rfl in + let ⟨k, hik, hk⟩ := hy i _ rfl in + let ⟨l, hl⟩ := directed_order.directed j k in + ⟨l, le_trans hij hl.1, begin + have hx := congr_arg (f j l hl.1) hj, + have hy := congr_arg (f k l hl.2) hk, + simp only [directed_system.Hcomp f] at hx hy, + rw [← hxy, linear_map.map_add, linear_map.map_add, hx, hy], + simp + end⟩) + (λ a x hx i _ hax, + let ⟨j, hij, hj⟩ := hx i _ rfl in + ⟨j, hij, begin + clear _let_match, + subst hax, + simp [linear_map.map_smul, hj] + end⟩) + _ _ + (begin + rw [sub_zero, direct_sum.to_module.of, totalize], + dsimp, + rw [dif_pos (le_refl _), directed_system.Hid f] + end) +#print subtype.val theorem of.zero_exact {i x} (H : of ι G R f i x = 0) : ∃ j hij, f i j hij x = (0 : G j) := -begin - rcases mem_span_iff_lc.1 ((submodule.quotient.eq _).1 H) with ⟨l, h1, h2⟩, rw sub_zero at h2, clear H, - revert l i x, intro l, haveI : decidable_eq R := classical.dec_eq R, - apply finsupp.induction l, - { intros i x h1 h2, rw [linear_map.map_zero, ← linear_map.map_zero (direct_sum.of ι G i)] at h2, - have := direct_sum.of_inj _ h2, subst this, - exact ⟨i, le_refl i, linear_map.map_zero _⟩ }, - intros a b f haf hb0 ih i x h1 h2, - sorry -end +@span_induction _ _ _ _ _ _ _ + (λ y, ∀ i (x : G i), (direct_sum.to_module ι (λ j, totalize G f j i)) y = x → + ∃ j hij, f i j hij x = (0 : G j)) + ((submodule.quotient.eq _).1 H) + (λ y hy', let ⟨i, j, Hij, z, hy⟩ := mem_thing.1 hy' in + (λ k xk hyk, + let ⟨l, hl⟩ := directed_order.directed j k in + ⟨l, hl.2, + --have i ≤ k → j ≤ k, from _, + begin + clear _let_match _let_match hy', + subst hyk, subst hy, + simp [direct_sum.apply_eq_component, direct_sum.ext_iff, linear_map.map_sub, + totalize_apply], + split_ifs; + simp [directed_system.Hcomp f] at *, + + + + + end⟩)) + (λ i x hix, ⟨i, le_refl _, + by rw [← linear_map.map_zero (direct_sum.of ι G i), direct_sum.to_module.of, + linear_map.map_zero] at hix; + rw [directed_system.Hid f, hix]⟩) + (λ x y hx hy i _ hxy, + let ⟨j, hij, hj⟩ := hx i _ rfl in + let ⟨k, hik, hk⟩ := hy i _ rfl in + let ⟨l, hl⟩ := directed_order.directed j k in + ⟨l, le_trans hij hl.1, begin + have hx := congr_arg (f j l hl.1) hj, + have hy := congr_arg (f k l hl.2) hk, + simp only [directed_system.Hcomp f] at hx hy, + rw [← hxy, linear_map.map_add, linear_map.map_add, hx, hy], + simp + end⟩) + (λ a x hx i _ hax, + let ⟨j, hij, hj⟩ := hx i _ rfl in + ⟨j, hij, begin + clear _let_match, + subst hax, + simp [linear_map.map_smul, hj] + end⟩) + _ _ + (begin + rw [sub_zero, direct_sum.to_module.of, totalize], + dsimp, + rw [dif_pos (le_refl _), directed_system.Hid f] + end) end module diff --git a/src/data/dfinsupp.lean b/src/data/dfinsupp.lean index a90bc76cf58dc..ddb973cbd3ee0 100644 --- a/src/data/dfinsupp.lean +++ b/src/data/dfinsupp.lean @@ -728,6 +728,13 @@ begin all_goals { intros, simp } end +lemma map_sum {R δ : Type*} [ring R] [add_comm_group δ] [module R δ] [add_comm_group γ] + [module R γ] [Π i, add_comm_monoid (β i)] + [Π i, decidable_pred (eq (0 : β i))] {f : Π₀ i, β i} {h : Π i, β i → γ} {l : γ →ₗ δ} : + l (f.sum h) = f.support.sum (λ x, (h x)) := + + + @[to_additive dfinsupp.sum_subtype_domain_index] lemma prod_subtype_domain_index [Π i, has_zero (β i)] [Π i, decidable_pred (eq (0 : β i))] [comm_monoid γ] {v : Π₀ i, β i} {p : ι → Prop} [decidable_pred p] diff --git a/src/linear_algebra/direct_sum_module.lean b/src/linear_algebra/direct_sum_module.lean index 78d2ffccc2e36..c875957bb6fa7 100644 --- a/src/linear_algebra/direct_sum_module.lean +++ b/src/linear_algebra/direct_sum_module.lean @@ -36,6 +36,11 @@ theorem mk_smul (s : finset ι) (c : R) (x) : mk β s (c • x) = c • mk β s theorem of_smul (i : ι) (c : R) (x) : of β i (c • x) = c • of β i x := (lof R ι β i).map_smul c x +lemma sum_of [Π i, decidable_pred (eq (0 : β i))] (f : direct_sum R ι β) : + f.sum (λ i, of ι β i) = f := +by dsimp [of, dfinsupp.lsingle]; unfold_coes; simp; + exact @dfinsupp.sum_single ι β _ _ _ f + variables {γ : Type u₁} [add_comm_group γ] [module R γ] variables (φ : Π i, β i →ₗ[R] γ) @@ -72,4 +77,29 @@ protected def lid (M : Type v) [add_comm_group M] [module R M] : { .. direct_sum.id M, .. to_module R punit M (λ i, linear_map.id) } +def component (ι : Type*) [decidable_eq ι] (β : ι → Type*) + [Π i, add_comm_group (β i)] [Π i, module R (β i)] + (i : ι) : direct_sum R ι β →ₗ β i := +{ to_fun := λ f, f i, + add := λ _ _, dfinsupp.add_apply, + smul := λ _ _, dfinsupp.smul_apply } + +lemma of_apply (i : ι) (b : β i) : ((of ι β i) b) i = b := +by rw [of, dfinsupp.lsingle_apply, dfinsupp.single_apply, dif_pos rfl] + +lemma apply_eq_component (f : direct_sum R ι β) (i : ι) : + f i = component ι β i f := rfl + +@[simp] lemma component.of (i : ι) (b : β i) : + component ι β i ((of ι β i) b) = b := +of_apply i b + +@[extensionality] lemma ext {f g : direct_sum R ι β} + (h : ∀ i, component ι β i f = component ι β i g) : f = g := +dfinsupp.ext h + +lemma ext_iff {f g : direct_sum R ι β} : f = g ↔ + ∀ i, component ι β i f = component ι β i g := +⟨λ h _, by rw h, ext⟩ + end direct_sum From 9ad20fcf89a6485c3460c7f62d7d08e92bdc4be4 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Fri, 1 Feb 2019 12:36:27 +0000 Subject: [PATCH 05/22] prove of_zero.exact --- src/algebra/direct_limit.lean | 204 ++++++---------------- src/data/dfinsupp.lean | 14 +- src/linear_algebra/direct_sum_module.lean | 11 +- 3 files changed, 67 insertions(+), 162 deletions(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 43adc8ede9123..6f51374d60d36 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -2,17 +2,6 @@ import linear_algebra.direct_sum_module linear_algebra.linear_combination tactic universes u v w u₁ -namespace tactic.interactive -open tactic -meta def clear_aux_decl : tactic unit := -do l ← local_context, -match l with -| [] := skip -| (e::l) := cond e.is_aux_decl (tactic.clear e) skip -end - -end tactic.interactive - open lattice submodule class directed_order (α : Type u) extends partial_order α := @@ -121,154 +110,61 @@ if h : i ≤ j else by dsimp [totalize]; rw [dif_neg h, dif_neg h, linear_map.zero_apply] -- set_option trace.simplify.rewrite true +lemma to_module_totalize_of_le {x : direct_sum R ι G} {i j : ι} + (hij : i ≤ j) (hx : ∀ k ∈ x.support, k ≤ i) : + direct_sum.to_module ι (λ k, totalize G f k j) x = + f i j hij (direct_sum.to_module ι (λ k, totalize G f k i) x) := +begin + rw [← @dfinsupp.sum_single ι G _ _ _ x, dfinsupp.sum], + simp only [linear_map.map_sum], + refine finset.sum_congr rfl (λ k hk, _), + rw [direct_sum.single_eq_of], + simp [totalize_apply, hx k hk, le_trans (hx k hk) hij, + directed_system.Hcomp f] +end - -theorem of.zero_exact {i x} (H : of ι G R f i x = 0) : - ∃ j hij, f i j hij x = (0 : G j) := -@span_induction _ _ _ _ _ _ _ - (λ y, ∀ i (x : G i), y = direct_sum.of ι G i x → - ∃ j hij, f i j hij x = (0 : G j)) - ((submodule.quotient.eq _).1 H) - (λ y hy', let ⟨i, j, Hij, z, hy⟩ := mem_thing.1 hy' in - λ k xk hyk, - let ⟨l, hl⟩ := directed_order.directed i k in - ⟨l, hl.2, begin - clear_aux_decl, - subst hyk, - have := congr_arg (direct_sum.component ι G k) hy, - clear hy', - simp at this, - subst this, - - - +lemma of.zero_exact_aux {x : direct_sum R ι G} (H : x ∈ span (thing G f)) : + ∃ j, (∀ k ∈ x.support, k ≤ j) ∧ direct_sum.to_module ι (λ i, totalize G f i j) x = (0 : G j) := +span_induction H + (λ x hx, let ⟨i, j, hij, y, hxy⟩ := mem_thing.1 hx in + let ⟨k, hik, hjk⟩ := directed_order.directed i j in + ⟨k, begin + clear_, + subst hxy, + simp [linear_map.map_sub, totalize_apply, hik, hjk, + directed_system.Hcomp f, direct_sum.apply_eq_component, + direct_sum.component.of], + intros, + split_ifs at *; + simp * at * + end⟩) + ⟨default ι, λ _ h, (finset.not_mem_empty _ h).elim, linear_map.map_zero _⟩ + (λ x y ⟨i, hi, hxi⟩ ⟨j, hj, hyj⟩, + let ⟨k, hik, hjk⟩ := directed_order.directed i j in + ⟨k, λ l hl, + (finset.mem_union.1 (dfinsupp.support_add hl)).elim + (λ hl, le_trans (hi _ hl) hik) + (λ hl, le_trans (hj _ hl) hjk), + begin + clear_, + rw [← @dfinsupp.sum_single ι G _ _ _ x], + simp [linear_map.map_add, hxi, hyj, + to_module_totalize_of_le G f hik hi, + to_module_totalize_of_le G f hjk hj] end⟩) - _ - begin - simp, - - end - _ - _ _ (sub_zero _) + (λ a x ⟨i, hi, hxi⟩, + ⟨i, λ k hk, hi k (dfinsupp.support_smul hk), + by simp [linear_map.map_smul, hxi]⟩) -theorem of.zero_exact {i x} (H : of ι G R f i x = 0) : - ∃ j hij, f i j hij x = (0 : G j) := -@span_induction _ _ _ _ _ _ _ - (λ y, ∀ i (x : G i), (direct_sum.to_module ι (λ j, totalize G f j i)) y = x → - ∃ j hij, f i j hij x = (0 : G j)) - ((submodule.quotient.eq _).1 H) - (λ y hy', let ⟨i, j, Hij, z, hy⟩ := mem_thing.1 hy' in - (λ k xk hyk, - let ⟨l, hl⟩ := directed_order.directed j k in - ⟨l, hl.2, begin - clear _let_match _let_match, - subst hyk, - have := congr_arg ((direct_sum.of ι G l).comp (direct_sum.to_module ι (λ i, totalize G f i l))) hy, - simp only [direct_sum.to_module.of, linear_map.map_sub] at this, - simp [totalize_apply] at *, - rw [dif_pos (le_trans Hij hl.1), dif_pos hl.1, directed_system.Hcomp f, - add_neg_self] at this, - clear H hy' x hy Hij z i i, - rw [← direct_sum.sum_of y, dfinsupp.sum] at *, - refine direct_sum.of_inj l _, - simp only [linear_map.map_sum, direct_sum.to_module.of, - direct_sum.ext_iff, linear_map.map_zero] at *, - assume i, - calc finset.sum (dfinsupp.support y) - (λ (m : ι), direct_sum.component ι G i - ((direct_sum.of ι G l) ((f k l _) ((totalize G f m k) (direct_sum.component ι G m y))))) - = direct_sum.component ι G i - ((direct_sum.of ι G l) ((f k l _) ((totalize G f l k) (direct_sum.component ι G l y)))) : - finset.sum_eq_single _ begin - intros, - simp [direct_sum.apply_eq_component, totalize_apply], - - end _ - ... = _ : sorry - - - - end⟩)) - (λ i x hix, ⟨i, le_refl _, - by rw [← linear_map.map_zero (direct_sum.of ι G i), direct_sum.to_module.of, - linear_map.map_zero] at hix; - rw [directed_system.Hid f, hix]⟩) - (λ x y hx hy i _ hxy, - let ⟨j, hij, hj⟩ := hx i _ rfl in - let ⟨k, hik, hk⟩ := hy i _ rfl in - let ⟨l, hl⟩ := directed_order.directed j k in - ⟨l, le_trans hij hl.1, begin - have hx := congr_arg (f j l hl.1) hj, - have hy := congr_arg (f k l hl.2) hk, - simp only [directed_system.Hcomp f] at hx hy, - rw [← hxy, linear_map.map_add, linear_map.map_add, hx, hy], - simp - end⟩) - (λ a x hx i _ hax, - let ⟨j, hij, hj⟩ := hx i _ rfl in - ⟨j, hij, begin - clear _let_match, - subst hax, - simp [linear_map.map_smul, hj] - end⟩) - _ _ - (begin - rw [sub_zero, direct_sum.to_module.of, totalize], - dsimp, - rw [dif_pos (le_refl _), directed_system.Hid f] - end) -#print subtype.val theorem of.zero_exact {i x} (H : of ι G R f i x = 0) : ∃ j hij, f i j hij x = (0 : G j) := -@span_induction _ _ _ _ _ _ _ - (λ y, ∀ i (x : G i), (direct_sum.to_module ι (λ j, totalize G f j i)) y = x → - ∃ j hij, f i j hij x = (0 : G j)) - ((submodule.quotient.eq _).1 H) - (λ y hy', let ⟨i, j, Hij, z, hy⟩ := mem_thing.1 hy' in - (λ k xk hyk, - let ⟨l, hl⟩ := directed_order.directed j k in - ⟨l, hl.2, - --have i ≤ k → j ≤ k, from _, - begin - clear _let_match _let_match hy', - subst hyk, subst hy, - simp [direct_sum.apply_eq_component, direct_sum.ext_iff, linear_map.map_sub, - totalize_apply], - split_ifs; - simp [directed_system.Hcomp f] at *, - - - - - end⟩)) - (λ i x hix, ⟨i, le_refl _, - by rw [← linear_map.map_zero (direct_sum.of ι G i), direct_sum.to_module.of, - linear_map.map_zero] at hix; - rw [directed_system.Hid f, hix]⟩) - (λ x y hx hy i _ hxy, - let ⟨j, hij, hj⟩ := hx i _ rfl in - let ⟨k, hik, hk⟩ := hy i _ rfl in - let ⟨l, hl⟩ := directed_order.directed j k in - ⟨l, le_trans hij hl.1, begin - have hx := congr_arg (f j l hl.1) hj, - have hy := congr_arg (f k l hl.2) hk, - simp only [directed_system.Hcomp f] at hx hy, - rw [← hxy, linear_map.map_add, linear_map.map_add, hx, hy], - simp - end⟩) - (λ a x hx i _ hax, - let ⟨j, hij, hj⟩ := hx i _ rfl in - ⟨j, hij, begin - clear _let_match, - subst hax, - simp [linear_map.map_smul, hj] - end⟩) - _ _ - (begin - rw [sub_zero, direct_sum.to_module.of, totalize], - dsimp, - rw [dif_pos (le_refl _), directed_system.Hid f] - end) +let ⟨j, hj, hxj⟩ := of.zero_exact_aux G f + ((submodule.quotient.eq _).1 H) in +if hx0 : x = 0 then ⟨i, le_refl _, by simp [hx0]⟩ +else + have hij : i ≤ j, from hj _ $ + by simp [direct_sum.apply_eq_component, hx0], + ⟨j, hij, by simpa [totalize_apply, hij] using hxj⟩ end module diff --git a/src/data/dfinsupp.lean b/src/data/dfinsupp.lean index ddb973cbd3ee0..baa0bb4f4df11 100644 --- a/src/data/dfinsupp.lean +++ b/src/data/dfinsupp.lean @@ -559,6 +559,13 @@ support_zip_with support (-f) = support f := by ext i; simp +local attribute [instance] dfinsupp.to_module + +lemma support_smul {γ : Type w} [ring γ] [Π i, add_comm_group (β i)] [Π i, module γ (β i)] + [Π (i : ι), decidable_pred (eq (0 : β i))] + {b : γ} {v : Π₀ i, β i} : (b • v).support ⊆ v.support := +λ x, by simp [dfinsupp.mem_support_iff, not_imp_not] {contextual := tt} + instance [decidable_eq ι] [Π i, has_zero (β i)] [Π i, decidable_eq (β i)] : decidable_eq (Π₀ i, β i) := assume f g, decidable_of_iff (f.support = g.support ∧ (∀i∈f.support, f i = g i)) ⟨assume ⟨h₁, h₂⟩, ext $ assume i, @@ -728,13 +735,6 @@ begin all_goals { intros, simp } end -lemma map_sum {R δ : Type*} [ring R] [add_comm_group δ] [module R δ] [add_comm_group γ] - [module R γ] [Π i, add_comm_monoid (β i)] - [Π i, decidable_pred (eq (0 : β i))] {f : Π₀ i, β i} {h : Π i, β i → γ} {l : γ →ₗ δ} : - l (f.sum h) = f.support.sum (λ x, (h x)) := - - - @[to_additive dfinsupp.sum_subtype_domain_index] lemma prod_subtype_domain_index [Π i, has_zero (β i)] [Π i, decidable_pred (eq (0 : β i))] [comm_monoid γ] {v : Π₀ i, β i} {p : ι → Prop} [decidable_pred p] diff --git a/src/linear_algebra/direct_sum_module.lean b/src/linear_algebra/direct_sum_module.lean index c875957bb6fa7..e8e3c965acddd 100644 --- a/src/linear_algebra/direct_sum_module.lean +++ b/src/linear_algebra/direct_sum_module.lean @@ -25,6 +25,10 @@ instance : module R (direct_sum ι β) := dfinsupp.to_module variables R ι β def lmk : Π s : finset ι, (Π i : (↑s : set ι), β i.1) →ₗ[R] direct_sum ι β := dfinsupp.lmk β R +lemma single_eq_of (i : ι) (b : β i) : + dfinsupp.single i b = of ι β i b := rfl + +variables {ι β} def lof : Π i : ι, β i →ₗ[R] direct_sum ι β := dfinsupp.lsingle β R @@ -90,10 +94,15 @@ by rw [of, dfinsupp.lsingle_apply, dfinsupp.single_apply, dif_pos rfl] lemma apply_eq_component (f : direct_sum R ι β) (i : ι) : f i = component ι β i f := rfl -@[simp] lemma component.of (i : ι) (b : β i) : +@[simp] lemma component.of_self (i : ι) (b : β i) : component ι β i ((of ι β i) b) = b := of_apply i b +lemma component.of (i j : ι) (b : β j) : + component ι β i ((of ι β j) b) = + if h : j = i then eq.rec_on h b else 0 := +dfinsupp.single_apply + @[extensionality] lemma ext {f g : direct_sum R ι β} (h : ∀ i, component ι β i f = component ι β i g) : f = g := dfinsupp.ext h From 6ddf081f61afe4fc499ff28d850fd2704849a55e Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Fri, 1 Feb 2019 12:46:49 +0000 Subject: [PATCH 06/22] remove silly rewrite --- src/algebra/direct_limit.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 6f51374d60d36..bc3b07226e696 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -147,7 +147,6 @@ span_induction H (λ hl, le_trans (hj _ hl) hjk), begin clear_, - rw [← @dfinsupp.sum_single ι G _ _ _ x], simp [linear_map.map_add, hxi, hyj, to_module_totalize_of_le G f hik hi, to_module_totalize_of_le G f hjk hj] From 84ada9fcb3a079e8cf23ebbd3b0371438bed8ca6 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Fri, 1 Feb 2019 17:45:16 +0000 Subject: [PATCH 07/22] slightly shorten proof --- src/algebra/direct_limit.lean | 7 ++----- 1 file changed, 2 insertions(+), 5 deletions(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index bc3b07226e696..e14a794c30e6c 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -145,12 +145,9 @@ span_induction H (finset.mem_union.1 (dfinsupp.support_add hl)).elim (λ hl, le_trans (hi _ hl) hik) (λ hl, le_trans (hj _ hl) hjk), - begin - clear_, - simp [linear_map.map_add, hxi, hyj, + by simp [linear_map.map_add, hxi, hyj, to_module_totalize_of_le G f hik hi, - to_module_totalize_of_le G f hjk hj] - end⟩) + to_module_totalize_of_le G f hjk hj]⟩) (λ a x ⟨i, hi, hxi⟩, ⟨i, λ k hk, hi k (dfinsupp.support_smul hk), by simp [linear_map.map_smul, hxi]⟩) From c752316a40948c00d7ba488cd2aa370e0c3af500 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Wed, 20 Feb 2019 08:01:04 +0000 Subject: [PATCH 08/22] direct limit of modules --- src/algebra/direct_limit.lean | 72 ++++++++++++++++++++++++++++++++++- 1 file changed, 70 insertions(+), 2 deletions(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index e14a794c30e6c..66217a22636f3 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -11,6 +11,8 @@ variables {ι : Type u} [inhabited ι] variables [directed_order ι] [decidable_eq ι] variables (G : ι → Type v) [Π i, decidable_eq (G i)] +namespace module + class directed_system {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] (f : Π i j, i ≤ j → G i →ₗ G j) : Prop := (Hid : ∀ i x, f i i (le_refl i) x = x) @@ -34,8 +36,6 @@ lemma mem_thing {G : ι → Type v} [Π i, decidable_eq (G i)] namespace direct_limit -section module - variables {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] variables (f : Π i j, i ≤ j → G i →ₗ G j) [directed_system G f] include R @@ -161,7 +161,75 @@ else have hij : i ≤ j, from hj _ $ by simp [direct_sum.apply_eq_component, hx0], ⟨j, hij, by simpa [totalize_apply, hij] using hxj⟩ +end direct_limit + +@[extensionality] theorem ext {α : Type u} [ring α] {β : Type v} [add_comm_group β] : + Π (M₁ M₂ : module α β) + (H : ∀ r x, @has_scalar.smul α β M₁.to_has_scalar r x = @has_scalar.smul α β M₂.to_has_scalar r x), + M₁ = M₂ := +by rintro ⟨⟨⟨f⟩, _, _, _, _, _, _⟩⟩ ⟨⟨⟨g⟩, _, _, _, _, _, _⟩⟩ H; congr' 3; ext; apply H end module + +namespace add_comm_group + +variables {α : Type u} [add_comm_group α] + +protected def module : module ℤ α := module.of_core +{ smul := gsmul, + mul_smul := λ r s x, gsmul_mul x r s, + smul_add := λ r x y, gsmul_add x y r, + add_smul := λ r s x, add_gsmul x r s, + one_smul := one_gsmul } + +instance module.subsingleton : subsingleton (module ℤ α) := +begin + constructor, intros M₁ M₂, ext i x, refine int.induction_on i _ _ _, + { rw [zero_smul, zero_smul] }, + { intros i ih, rw [add_smul, add_smul, ih, one_smul, one_smul] }, + { intros i ih, rw [sub_smul, sub_smul, ih, one_smul, one_smul] } +end + +end add_comm_group + +local attribute [instance] add_comm_group.module + +namespace is_add_group_hom + +variables {α : Type u} {β : Type v} [add_comm_group α] [add_comm_group β] + +def to_linear_map (f : α → β) [is_add_group_hom f] : α →ₗ β := +{ to_fun := f, + add := is_add_group_hom.add f, + smul := λ i x, int.induction_on i (by rw [zero_smul, zero_smul, is_add_group_hom.zero f]) + (λ i ih, by rw [add_smul, add_smul, is_add_group_hom.add f, ih, one_smul, one_smul]) + (λ i ih, by rw [sub_smul, sub_smul, is_add_group_hom.sub f, ih, one_smul, one_smul]) } + +end is_add_group_hom + +namespace add_comm_group + +class directed_system [Π i, add_comm_group (G i)] (f : Π i j, i ≤ j → G i → G j) : Prop := +(Hid : ∀ i x, f i i (le_refl i) x = x) +(Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) + +def direct_limit [Π i, add_comm_group (G i)] + (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] : Type* := +@module.direct_limit ι _ _ _ G _ ℤ _ _ _ + (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) + ⟨directed_system.Hid f, directed_system.Hcomp f⟩ + +namespace direct_limit + +variables [Π i, add_comm_group (G i)] (f : Π i j, i ≤ j → G i → G j) +variables [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] + +instance : add_comm_group (direct_limit G f) := +@module.direct_limit.add_comm_group _ _ _ _ _ _ _ _ _ _ + (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) + ⟨directed_system.Hid f, directed_system.Hcomp f⟩ + end direct_limit + +end add_comm_group From 88caa1c0afc1645b39243cb7d136d216034b67c1 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Wed, 20 Feb 2019 09:24:29 +0000 Subject: [PATCH 09/22] upgrade mathlib --- src/algebra/direct_limit.lean | 118 +++++++++++----------- src/linear_algebra/direct_sum_module.lean | 55 +++++----- 2 files changed, 85 insertions(+), 88 deletions(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 66217a22636f3..18ea0a75993d1 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -7,71 +7,67 @@ open lattice submodule class directed_order (α : Type u) extends partial_order α := (directed : ∀ i j : α, ∃ k, i ≤ k ∧ j ≤ k) -variables {ι : Type u} [inhabited ι] +variables {R : Type u} [ring R] +variables {ι : Type v} [inhabited ι] variables [directed_order ι] [decidable_eq ι] -variables (G : ι → Type v) [Π i, decidable_eq (G i)] +variables (G : ι → Type w) [Π i, decidable_eq (G i)] [Π i, add_comm_group (G i)] namespace module -class directed_system {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] - (f : Π i j, i ≤ j → G i →ₗ G j) : Prop := +variables [Π i, module R (G i)] + +class directed_system (f : Π i j, i ≤ j → G i →ₗ[R] G j) : Prop := (Hid : ∀ i x, f i i (le_refl i) x = x) (Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) -def thing {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] - (f : Π i j, i ≤ j → G i →ₗ G j) [directed_system G f] : set (direct_sum R ι G) := - ⋃ i j (H : i ≤ j) x, - ({ direct_sum.of ι G i x - direct_sum.of ι G j (f i j H x) } : set $ direct_sum R ι G) +variables (f : Π i j, i ≤ j → G i →ₗ[R] G j) [directed_system G f] -def direct_limit {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] - (f : Π i j, i ≤ j → G i →ₗ G j) [directed_system G f] : Type (max u v) := -quotient $ span $ thing G f +def thing : set (direct_sum ι G) := +⋃ i j (H : i ≤ j) x, { direct_sum.lof R ι G i x - direct_sum.lof R ι G j (f i j H x) } -lemma mem_thing {G : ι → Type v} [Π i, decidable_eq (G i)] - {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] - {f : Π i j, i ≤ j → G i →ₗ G j} [directed_system G f] {a : direct_sum R ι G} : - a ∈ thing G f ↔ ∃ (i j) (H : i ≤ j) x, - a = direct_sum.of ι G i x - direct_sum.of ι G j (f i j H x) := - by simp [thing, set.mem_Union] +def direct_limit : Type (max v w) := +(span R $ thing G f).quotient -namespace direct_limit +variables {G f} +lemma mem_thing {a : direct_sum ι G} : a ∈ thing G f ↔ ∃ (i j) (H : i ≤ j) x, + a = direct_sum.lof R ι G i x - direct_sum.lof R ι G j (f i j H x) := +by simp only [thing, set.mem_Union, set.mem_singleton_iff] +variables (G f) -variables {R : Type w} [comm_ring R] [Π i, add_comm_group (G i)] [Π i, module R (G i)] -variables (f : Π i j, i ≤ j → G i →ₗ G j) [directed_system G f] -include R +namespace direct_limit instance : add_comm_group (direct_limit G f) := quotient.add_comm_group _ instance : module R (direct_limit G f) := quotient.module _ variables (ι R) -def of (i) : G i →ₗ direct_limit G f := -(mkq _).comp $ direct_sum.of ι G i +def of (i) : G i →ₗ[R] direct_limit G f := +(mkq _).comp $ direct_sum.lof R ι G i variables {ι R} -theorem of_f {i j hij x} : (of ι G R f j (f i j hij x)) = of ι G R f i x := +theorem of_f {i j hij x} : (of R ι G f j (f i j hij x)) = of R ι G f i x := eq.symm $ (submodule.quotient.eq _).2 $ subset_span $ set.mem_Union.2 ⟨i, set.mem_Union.2 ⟨j, set.mem_Union.2 ⟨hij, set.mem_Union.2 ⟨x, or.inl rfl⟩⟩⟩⟩ -variables {P : Type u₁} [add_comm_group P] [module R P] (g : Π i, G i →ₗ P) +variables {P : Type u₁} [add_comm_group P] [module R P] (g : Π i, G i →ₗ[R] P) variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) include Hg variables (ι R) -def rec : direct_limit G f →ₗ P := -liftq _ (direct_sum.to_module _ g) +def rec : direct_limit G f →ₗ[R] P := +liftq _ (direct_sum.to_module R ι P g) (span_le.2 $ set.Union_subset $ λ i, set.Union_subset $ λ j, set.Union_subset $ λ hij, set.Union_subset $ λ x, set.singleton_subset_iff.2 $ linear_map.sub_mem_ker_iff.2 $ - by rw [direct_sum.to_module.of, direct_sum.to_module.of, Hg]) + by rw [direct_sum.to_module_lof, direct_sum.to_module_lof, Hg]) variables {ι R} omit Hg -lemma rec_of {i} (x) : rec ι G R f g Hg (of ι G R f i x) = g i x := -direct_sum.to_module.of _ _ +lemma rec_of {i} (x) : rec R ι G f g Hg (of R ι G f i x) = g i x := +direct_sum.to_module_lof R _ _ -theorem rec_unique (F : direct_limit G f →ₗ P) (x) : - F x = rec ι G R f (λ i, F.comp $ of ι G R f i) +theorem rec_unique (F : direct_limit G f →ₗ[R] P) (x) : + F x = rec R ι G f (λ i, F.comp $ of R ι G f i) (λ i j hij x, by rw [linear_map.comp_apply, of_f]; refl) x := quotient.induction_on' x $ λ y, direct_sum.induction_on y ((linear_map.map_zero _).trans (linear_map.map_zero _).symm) @@ -79,7 +75,7 @@ quotient.induction_on' x $ λ y, direct_sum.induction_on y (λ x y ihx ihy, (linear_map.map_add F (quotient.mk' x) (quotient.mk' y)).trans $ ihx.symm ▸ ihy.symm ▸ (linear_map.map_add _ _ _).symm) -theorem exists_of (z : direct_limit G f) : ∃ i x, z = of ι G R f i x := +theorem exists_of (z : direct_limit G f) : ∃ i x, z = of R ι G f i x := quotient.induction_on' z $ λ z, direct_sum.induction_on z ⟨default _, 0, (linear_map.map_zero _).symm⟩ (λ i x, ⟨i, x, rfl⟩) @@ -87,56 +83,59 @@ quotient.induction_on' z $ λ z, direct_sum.induction_on z ⟨k, f i k hik x + f j k hjk y, by rw [linear_map.map_add, of_f, of_f, ← ihx, ← ihy]; refl⟩) @[elab_as_eliminator] theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit G f) - (ih : ∀ i x, C (of ι G R f i x)) : C z := + (ih : ∀ i x, C (of R ι G f i x)) : C z := let ⟨i, x, h⟩ := exists_of G f z in h.symm ▸ ih i x lemma span_preimage_le_comap_span {R M N: Type*} [ring R] [add_comm_group M] [module R M] - [add_comm_group N] [module R N] (f : N →ₗ M) (s : set M) : span (f ⁻¹' s) ≤ (span s).comap f := + [add_comm_group N] [module R N] (f : N →ₗ[R] M) (s : set M) : span R (f ⁻¹' s) ≤ (span R s).comap f := λ x h, span_induction h (by simp only [set.preimage, set.mem_set_of_eq, mem_comap]; exact λ x h, subset_span h) - (zero_mem ((span s).comap f)) - (λ _ _ hx hy, add_mem ((span s).comap f) hx hy) - (λ _ _ h, smul_mem ((span s).comap f) _ h) + (zero_mem ((span R s).comap f)) + (λ _ _ hx hy, add_mem ((span R s).comap f) hx hy) + (λ _ _ h, smul_mem ((span R s).comap f) _ h) +section totalize local attribute [instance, priority 0] classical.dec -noncomputable def totalize : Π i j, G i →ₗ G j := +noncomputable def totalize : Π i j, G i →ₗ[R] G j := λ i j, if h : i ≤ j then f i j h else 0 lemma totalize_apply (i j x) : totalize G f i j x = if h : i ≤ j then f i j h x else 0 := if h : i ≤ j - then by dsimp [totalize]; rw [dif_pos h, dif_pos h] - else by dsimp [totalize]; rw [dif_neg h, dif_neg h, linear_map.zero_apply] --- set_option trace.simplify.rewrite true + then by dsimp only [totalize]; rw [dif_pos h, dif_pos h] + else by dsimp only [totalize]; rw [dif_neg h, dif_neg h, linear_map.zero_apply] +end totalize -lemma to_module_totalize_of_le {x : direct_sum R ι G} {i j : ι} +lemma to_module_totalize_of_le {x : direct_sum ι G} {i j : ι} (hij : i ≤ j) (hx : ∀ k ∈ x.support, k ≤ i) : - direct_sum.to_module ι (λ k, totalize G f k j) x = - f i j hij (direct_sum.to_module ι (λ k, totalize G f k i) x) := + direct_sum.to_module R ι (G j) (λ k, totalize G f k j) x = + f i j hij (direct_sum.to_module R ι (G i) (λ k, totalize G f k i) x) := begin rw [← @dfinsupp.sum_single ι G _ _ _ x, dfinsupp.sum], simp only [linear_map.map_sum], refine finset.sum_congr rfl (λ k hk, _), - rw [direct_sum.single_eq_of], - simp [totalize_apply, hx k hk, le_trans (hx k hk) hij, - directed_system.Hcomp f] + rw direct_sum.single_eq_lof R k (x k), + simp [totalize_apply, hx k hk, le_trans (hx k hk) hij, directed_system.Hcomp f] end -lemma of.zero_exact_aux {x : direct_sum R ι G} (H : x ∈ span (thing G f)) : - ∃ j, (∀ k ∈ x.support, k ≤ j) ∧ direct_sum.to_module ι (λ i, totalize G f i j) x = (0 : G j) := +lemma of.zero_exact_aux {x : direct_sum ι G} (H : x ∈ span R (thing G f)) : + ∃ j, (∀ k ∈ x.support, k ≤ j) ∧ direct_sum.to_module R ι (G j) (λ i, totalize G f i j) x = (0 : G j) := span_induction H (λ x hx, let ⟨i, j, hij, y, hxy⟩ := mem_thing.1 hx in let ⟨k, hik, hjk⟩ := directed_order.directed i j in ⟨k, begin clear_, subst hxy, + split, + { intros i0 hi0, + rw [dfinsupp.mem_support_iff, dfinsupp.sub_apply, ← direct_sum.single_eq_lof, + ← direct_sum.single_eq_lof, dfinsupp.single_apply, dfinsupp.single_apply] at hi0, + split_ifs at hi0 with hi hj hj, { rwa hi at hik }, { rwa hi at hik }, { rwa hj at hjk }, + exfalso, apply hi0, rw sub_zero }, simp [linear_map.map_sub, totalize_apply, hik, hjk, directed_system.Hcomp f, direct_sum.apply_eq_component, direct_sum.component.of], - intros, - split_ifs at *; - simp * at * end⟩) ⟨default ι, λ _ h, (finset.not_mem_empty _ h).elim, linear_map.map_zero _⟩ (λ x y ⟨i, hi, hxi⟩ ⟨j, hj, hyj⟩, @@ -152,7 +151,7 @@ span_induction H ⟨i, λ k hk, hi k (dfinsupp.support_smul hk), by simp [linear_map.map_smul, hxi]⟩) -theorem of.zero_exact {i x} (H : of ι G R f i x = 0) : +theorem of.zero_exact {i x} (H : of R ι G f i x = 0) : ∃ j hij, f i j hij x = (0 : G j) := let ⟨j, hj, hxj⟩ := of.zero_exact_aux G f ((submodule.quotient.eq _).1 H) in @@ -171,7 +170,6 @@ by rintro ⟨⟨⟨f⟩, _, _, _, _, _, _⟩⟩ ⟨⟨⟨g⟩, _, _, _, _, _, _ end module - namespace add_comm_group variables {α : Type u} [add_comm_group α] @@ -199,7 +197,7 @@ namespace is_add_group_hom variables {α : Type u} {β : Type v} [add_comm_group α] [add_comm_group β] -def to_linear_map (f : α → β) [is_add_group_hom f] : α →ₗ β := +def to_linear_map (f : α → β) [is_add_group_hom f] : α →ₗ[ℤ] β := { to_fun := f, add := is_add_group_hom.add f, smul := λ i x, int.induction_on i (by rw [zero_smul, zero_smul, is_add_group_hom.zero f]) @@ -210,19 +208,19 @@ end is_add_group_hom namespace add_comm_group -class directed_system [Π i, add_comm_group (G i)] (f : Π i j, i ≤ j → G i → G j) : Prop := +class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := (Hid : ∀ i x, f i i (le_refl i) x = x) (Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) -def direct_limit [Π i, add_comm_group (G i)] +def direct_limit (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] : Type* := -@module.direct_limit ι _ _ _ G _ ℤ _ _ _ +@module.direct_limit ℤ _ ι _ _ _ G _ _ _ (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) ⟨directed_system.Hid f, directed_system.Hcomp f⟩ namespace direct_limit -variables [Π i, add_comm_group (G i)] (f : Π i j, i ≤ j → G i → G j) +variables (f : Π i j, i ≤ j → G i → G j) variables [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] instance : add_comm_group (direct_limit G f) := diff --git a/src/linear_algebra/direct_sum_module.lean b/src/linear_algebra/direct_sum_module.lean index e8e3c965acddd..d9afd5044c742 100644 --- a/src/linear_algebra/direct_sum_module.lean +++ b/src/linear_algebra/direct_sum_module.lean @@ -11,7 +11,7 @@ import linear_algebra.basic universes u v w u₁ -variables (R : Type u) [comm_ring R] +variables (R : Type u) [ring R] variables (ι : Type v) [decidable_eq ι] (β : ι → Type w) variables [Π i, add_comm_group (β i)] [Π i, module R (β i)] include R @@ -25,14 +25,13 @@ instance : module R (direct_sum ι β) := dfinsupp.to_module variables R ι β def lmk : Π s : finset ι, (Π i : (↑s : set ι), β i.1) →ₗ[R] direct_sum ι β := dfinsupp.lmk β R -lemma single_eq_of (i : ι) (b : β i) : - dfinsupp.single i b = of ι β i b := rfl - -variables {ι β} def lof : Π i : ι, β i →ₗ[R] direct_sum ι β := dfinsupp.lsingle β R -variables {R ι β} +variables {ι β} + +lemma single_eq_lof (i : ι) (b : β i) : + dfinsupp.single i b = lof R ι β i b := rfl theorem mk_smul (s : finset ι) (c : R) (x) : mk β s (c • x) = c • mk β s x := (lmk R ι β s).map_smul c x @@ -40,15 +39,15 @@ theorem mk_smul (s : finset ι) (c : R) (x) : mk β s (c • x) = c • mk β s theorem of_smul (i : ι) (c : R) (x) : of β i (c • x) = c • of β i x := (lof R ι β i).map_smul c x -lemma sum_of [Π i, decidable_pred (eq (0 : β i))] (f : direct_sum R ι β) : - f.sum (λ i, of ι β i) = f := -by dsimp [of, dfinsupp.lsingle]; unfold_coes; simp; +lemma sum_of [Π i, decidable_pred (eq (0 : β i))] (f : direct_sum ι β) : + f.sum (λ i, lof R ι β i) = f := +by dsimp [of, dfinsupp.lsingle]; unfold_coes; exact @dfinsupp.sum_single ι β _ _ _ f variables {γ : Type u₁} [add_comm_group γ] [module R γ] variables (φ : Π i, β i →ₗ[R] γ) -variables (R ι γ φ) +variables (ι γ φ) def to_module : direct_sum ι β →ₗ[R] γ := { to_fun := to_group (λ i, φ i), add := to_group_add _, @@ -56,7 +55,7 @@ def to_module : direct_sum ι β →ₗ[R] γ := (by rw [smul_zero, to_group_zero, smul_zero]) (λ i x, by rw [← of_smul, to_group_of, to_group_of, (φ i).map_smul c x]) (λ x y ihx ihy, by rw [smul_add, to_group_add, ihx, ihy, to_group_add, smul_add]) } -variables {R ι γ φ} +variables {ι γ φ} @[simp] lemma to_module_lof (i) (x : β i) : to_module R ι γ φ (lof R ι β i x) = φ i x := to_group_of (λ i, φ i) i x @@ -70,7 +69,7 @@ variables {ψ} {ψ' : direct_sum ι β →ₗ[R] γ} theorem to_module.ext (H : ∀ i, ψ.comp (lof R ι β i) = ψ'.comp (lof R ι β i)) (f : direct_sum ι β) : ψ f = ψ' f := -by rw [to_module.unique ψ, to_module.unique ψ', funext H] +by rw [to_module.unique R ψ, to_module.unique R ψ', funext H] def lset_to_set (S T : set ι) (H : S ⊆ T) : direct_sum S (β ∘ subtype.val) →ₗ direct_sum T (β ∘ subtype.val) := @@ -81,34 +80,34 @@ protected def lid (M : Type v) [add_comm_group M] [module R M] : { .. direct_sum.id M, .. to_module R punit M (λ i, linear_map.id) } -def component (ι : Type*) [decidable_eq ι] (β : ι → Type*) - [Π i, add_comm_group (β i)] [Π i, module R (β i)] - (i : ι) : direct_sum R ι β →ₗ β i := +variables (ι β) +def component (i : ι) : direct_sum ι β →ₗ[R] β i := { to_fun := λ f, f i, add := λ _ _, dfinsupp.add_apply, smul := λ _ _, dfinsupp.smul_apply } -lemma of_apply (i : ι) (b : β i) : ((of ι β i) b) i = b := -by rw [of, dfinsupp.lsingle_apply, dfinsupp.single_apply, dif_pos rfl] +variables {ι β} +@[simp] lemma lof_apply (i : ι) (b : β i) : ((lof R ι β i) b) i = b := +by rw [lof, dfinsupp.lsingle_apply, dfinsupp.single_apply, dif_pos rfl] -lemma apply_eq_component (f : direct_sum R ι β) (i : ι) : - f i = component ι β i f := rfl +lemma apply_eq_component (f : direct_sum ι β) (i : ι) : + f i = component R ι β i f := rfl -@[simp] lemma component.of_self (i : ι) (b : β i) : - component ι β i ((of ι β i) b) = b := -of_apply i b +@[simp] lemma component.lof_self (i : ι) (b : β i) : + component R ι β i ((lof R ι β i) b) = b := +lof_apply R i b lemma component.of (i j : ι) (b : β j) : - component ι β i ((of ι β j) b) = + component R ι β i ((lof R ι β j) b) = if h : j = i then eq.rec_on h b else 0 := dfinsupp.single_apply -@[extensionality] lemma ext {f g : direct_sum R ι β} - (h : ∀ i, component ι β i f = component ι β i g) : f = g := +@[extensionality] lemma ext {f g : direct_sum ι β} + (h : ∀ i, component R ι β i f = component R ι β i g) : f = g := dfinsupp.ext h -lemma ext_iff {f g : direct_sum R ι β} : f = g ↔ - ∀ i, component ι β i f = component ι β i g := -⟨λ h _, by rw h, ext⟩ +lemma ext_iff {f g : direct_sum ι β} : f = g ↔ + ∀ i, component R ι β i f = component R ι β i g := +⟨λ h _, by rw h, ext R⟩ end direct_sum From ce09a3c8294567845ff9eecc88d744d6aea880c1 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Wed, 20 Feb 2019 18:33:53 +0000 Subject: [PATCH 10/22] direct limit of rings --- src/algebra/direct_limit.lean | 120 +++++++++++++++++--- src/linear_algebra/tensor_product.lean | 6 +- src/ring_theory/free_comm_ring.lean | 145 +++++++++++++++++++++++++ 3 files changed, 252 insertions(+), 19 deletions(-) create mode 100644 src/ring_theory/free_comm_ring.lean diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 18ea0a75993d1..d9ec27f44b7bd 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -1,4 +1,7 @@ -import linear_algebra.direct_sum_module linear_algebra.linear_combination tactic.linarith +import linear_algebra.direct_sum_module +import linear_algebra.linear_combination +import ring_theory.free_comm_ring +import ring_theory.ideal_operations universes u v w u₁ @@ -10,11 +13,11 @@ class directed_order (α : Type u) extends partial_order α := variables {R : Type u} [ring R] variables {ι : Type v} [inhabited ι] variables [directed_order ι] [decidable_eq ι] -variables (G : ι → Type w) [Π i, decidable_eq (G i)] [Π i, add_comm_group (G i)] +variables (G : ι → Type w) [Π i, decidable_eq (G i)] namespace module -variables [Π i, module R (G i)] +variables [Π i, add_comm_group (G i)] [Π i, module R (G i)] class directed_system (f : Π i j, i ≤ j → G i →ₗ[R] G j) : Prop := (Hid : ∀ i x, f i i (le_refl i) x = x) @@ -23,15 +26,15 @@ class directed_system (f : Π i j, i ≤ j → G i →ₗ[R] G j) : Prop := variables (f : Π i j, i ≤ j → G i →ₗ[R] G j) [directed_system G f] def thing : set (direct_sum ι G) := -⋃ i j (H : i ≤ j) x, { direct_sum.lof R ι G i x - direct_sum.lof R ι G j (f i j H x) } +⋃ i j H, set.range $ λ x, direct_sum.lof R ι G i x - direct_sum.lof R ι G j (f i j H x) def direct_limit : Type (max v w) := (span R $ thing G f).quotient variables {G f} lemma mem_thing {a : direct_sum ι G} : a ∈ thing G f ↔ ∃ (i j) (H : i ≤ j) x, - a = direct_sum.lof R ι G i x - direct_sum.lof R ι G j (f i j H x) := -by simp only [thing, set.mem_Union, set.mem_singleton_iff] + direct_sum.lof R ι G i x - direct_sum.lof R ι G j (f i j H x) = a := +by simp only [thing, set.mem_Union, set.mem_range] variables (G f) namespace direct_limit @@ -39,28 +42,28 @@ namespace direct_limit instance : add_comm_group (direct_limit G f) := quotient.add_comm_group _ instance : module R (direct_limit G f) := quotient.module _ -variables (ι R) +variables (R ι) def of (i) : G i →ₗ[R] direct_limit G f := (mkq _).comp $ direct_sum.lof R ι G i -variables {ι R} +variables {R ι} theorem of_f {i j hij x} : (of R ι G f j (f i j hij x)) = of R ι G f i x := eq.symm $ (submodule.quotient.eq _).2 $ subset_span $ -set.mem_Union.2 ⟨i, set.mem_Union.2 ⟨j, set.mem_Union.2 ⟨hij, set.mem_Union.2 ⟨x, or.inl rfl⟩⟩⟩⟩ +set.mem_Union.2 ⟨i, set.mem_Union.2 ⟨j, set.mem_Union.2 ⟨hij, set.mem_range_self x⟩⟩⟩ variables {P : Type u₁} [add_comm_group P] [module R P] (g : Π i, G i →ₗ[R] P) variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) include Hg -variables (ι R) +variables (R ι) def rec : direct_limit G f →ₗ[R] P := liftq _ (direct_sum.to_module R ι P g) (span_le.2 $ set.Union_subset $ λ i, - set.Union_subset $ λ j, set.Union_subset $ λ hij, set.Union_subset $ λ x, - set.singleton_subset_iff.2 $ linear_map.sub_mem_ker_iff.2 $ + set.Union_subset $ λ j, set.Union_subset $ λ hij, + set.range_subset_iff.2 $ λ x, linear_map.sub_mem_ker_iff.2 $ by rw [direct_sum.to_module_lof, direct_sum.to_module_lof, Hg]) -variables {ι R} +variables {R ι} omit Hg lemma rec_of {i} (x) : rec R ι G f g Hg (of R ι G f i x) = g i x := @@ -170,6 +173,7 @@ by rintro ⟨⟨⟨f⟩, _, _, _, _, _, _⟩⟩ ⟨⟨⟨g⟩, _, _, _, _, _, _ end module + namespace add_comm_group variables {α : Type u} [add_comm_group α] @@ -208,6 +212,8 @@ end is_add_group_hom namespace add_comm_group +variables [Π i, add_comm_group (G i)] + class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := (Hid : ∀ i x, f i i (le_refl i) x = x) (Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) @@ -223,11 +229,93 @@ namespace direct_limit variables (f : Π i j, i ≤ j → G i → G j) variables [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] +def directed_system : module.directed_system G (λ (i j : ι) (hij : i ≤ j), is_add_group_hom.to_linear_map (f i j hij)) := +⟨directed_system.Hid f, directed_system.Hcomp f⟩ + +local attribute [instance] directed_system + instance : add_comm_group (direct_limit G f) := -@module.direct_limit.add_comm_group _ _ _ _ _ _ _ _ _ _ - (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) - ⟨directed_system.Hid f, directed_system.Hcomp f⟩ +module.direct_limit.add_comm_group G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) + +variables (P : Type u₁) [add_comm_group P] +variables (g : Π i, G i → P) [Π i, is_add_group_hom (g i)] +variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) + +set_option class.instance_max_depth 51 +def rec : direct_limit G f → P := +module.direct_limit.rec ℤ ι G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) + (λ i, is_add_group_hom.to_linear_map $ g i) Hg end direct_limit end add_comm_group + + +namespace ring + +variables [Π i, ring (G i)] +variables (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_ring_hom (f i j hij)] +variables [add_comm_group.directed_system G f] + +open free_comm_ring + +def thing1 : set (free_comm_ring Σ i, G i) := +⋃ i j H, set.range $ λ x, of ⟨j, f i j H x⟩ - of ⟨i, x⟩ + +def thing2 : set (free_comm_ring Σ i, G i) := +set.range $ λ i, of ⟨i, 1⟩ - 1 + +def thing3 : set (free_comm_ring Σ i, G i) := +⋃ i x, set.range $ λ y, of ⟨i, x + y⟩ - (of ⟨i, x⟩ + of ⟨i, y⟩) + +def thing4 : set (free_comm_ring Σ i, G i) := +⋃ i x, set.range $ λ y, of ⟨i, x * y⟩ - (of ⟨i, x⟩ * of ⟨i, y⟩) + +def direct_limit : Type (max v w) := +(ideal.span (thing1 G f ∪ thing2 G ∪ thing3 G ∪ thing4 G)).quotient + +namespace direct_limit + +instance : comm_ring (direct_limit G f) := +ideal.quotient.comm_ring _ + +def of (i) (x : G i) : direct_limit G f := +ideal.quotient.mk _ $ of ⟨i, x⟩ + +instance of.is_ring_hom (i) : is_ring_hom (of G f i) := +{ map_one := ideal.quotient.eq.2 $ subset_span $ or.inl $ or.inl $ or.inr $ set.mem_range_self i, + map_mul := λ x y, ideal.quotient.eq.2 $ subset_span $ or.inr $ set.mem_Union.2 ⟨i, + set.mem_Union.2 ⟨x, set.mem_range_self y⟩⟩, + map_add := λ x y, ideal.quotient.eq.2 $ subset_span $ or.inl $ or.inr $ set.mem_Union.2 ⟨i, + set.mem_Union.2 ⟨x, set.mem_range_self y⟩⟩ } + +variables (P : Type u₁) [comm_ring P] +variables (g : Π i, G i → P) [Π i, is_ring_hom (g i)] +variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) +include Hg + +open free_comm_ring + +def rec : direct_limit G f → P := +ideal.quotient.lift _ (free_comm_ring.lift $ λ x, g x.1 x.2) begin + suffices : ideal.span (thing1 G f ∪ thing2 G ∪ thing3 G ∪ thing4 G) ≤ + ideal.comap (free_comm_ring.lift (λ (x : Σ (i : ι), G i), g (x.fst) (x.snd))) ⊥, + { intros x hx, exact (mem_bot P).1 (this hx) }, + rw ideal.span_le, intros x hx, + simp only [thing1, thing2, thing3, thing4, set.mem_union, set.mem_Union, set.mem_range] at hx, + rw [mem_coe, ideal.mem_comap, mem_bot], + rcases hx with ⟨⟨i, j, hij, x, rfl⟩ | hx⟩, + { simp only [lift_sub, lift_of, Hg, sub_self] }, + { rcases hx with ⟨i, rfl⟩, simp only [lift_sub, lift_of, lift_one, is_ring_hom.map_one (g i), sub_self] }, + { rcases hx with ⟨i, x, y, rfl⟩, simp only [lift_sub, lift_of, lift_add, is_ring_hom.map_add (g i), sub_self] }, + { rcases hx with ⟨i, x, y, rfl⟩, simp only [lift_sub, lift_of, lift_mul, is_ring_hom.map_mul (g i), sub_self] } +end + +instance rec.is_ring_hom : is_ring_hom (rec G f P g Hg) := +⟨free_comm_ring.lift_one _, +λ x y, quotient.induction_on₂' x y $ λ p q, free_comm_ring.lift_mul _ _ _, +λ x y, quotient.induction_on₂' x y $ λ p q, free_comm_ring.lift_add _ _ _⟩ + +end direct_limit + +end ring diff --git a/src/linear_algebra/tensor_product.lean b/src/linear_algebra/tensor_product.lean index fbdd6882b843a..e6a92625602de 100644 --- a/src/linear_algebra/tensor_product.lean +++ b/src/linear_algebra/tensor_product.lean @@ -410,10 +410,10 @@ begin (lift $ direct_sum.to_module R _ _ $ λ i₁, flip $ direct_sum.to_module R _ _ $ λ i₂, flip $ curry $ direct_sum.lof R (ι₁ × ι₂) (λ i, β₁ i.1 ⊗[R] β₂ i.2) (i₁, i₂)) (direct_sum.to_module R _ _ $ λ i, map (direct_sum.lof R _ _ _) (direct_sum.lof R _ _ _)) - (linear_map.ext $ direct_sum.to_module.ext $ λ i, mk_compr₂_inj $ + (linear_map.ext $ direct_sum.to_module.ext _ $ λ i, mk_compr₂_inj $ linear_map.ext $ λ x₁, linear_map.ext $ λ x₂, _) - (mk_compr₂_inj $ linear_map.ext $ direct_sum.to_module.ext $ λ i₁, linear_map.ext $ λ x₁, - linear_map.ext $ direct_sum.to_module.ext $ λ i₂, linear_map.ext $ λ x₂, _); + (mk_compr₂_inj $ linear_map.ext $ direct_sum.to_module.ext _ $ λ i₁, linear_map.ext $ λ x₁, + linear_map.ext $ direct_sum.to_module.ext _ $ λ i₂, linear_map.ext $ λ x₂, _); repeat { rw compr₂_apply <|> rw comp_apply <|> rw id_apply <|> rw mk_apply <|> rw direct_sum.to_module_lof <|> rw map_tmul <|> rw lift.tmul <|> rw flip_apply <|> rw curry_apply }, diff --git a/src/ring_theory/free_comm_ring.lean b/src/ring_theory/free_comm_ring.lean new file mode 100644 index 0000000000000..20ff8528dbb8d --- /dev/null +++ b/src/ring_theory/free_comm_ring.lean @@ -0,0 +1,145 @@ +import group_theory.free_abelian_group + +universes u v + +def free_comm_ring (α : Type u) : Type u := +free_abelian_group $ multiset α + +namespace free_comm_ring + +variables (α : Type u) + +instance : add_comm_group (free_comm_ring α) := +{ .. free_abelian_group.add_comm_group (multiset α) } + +instance : semigroup (free_comm_ring α) := +{ mul := λ x, free_abelian_group.lift $ λ s2, free_abelian_group.lift (λ s1, free_abelian_group.of $ s1 + s2) x, + mul_assoc := λ x y z, begin + unfold has_mul.mul, + refine free_abelian_group.induction_on z rfl _ _ _, + { intros s3, rw [free_abelian_group.lift.of, free_abelian_group.lift.of], + refine free_abelian_group.induction_on y rfl _ _ _, + { intros s2, iterate 3 { rw free_abelian_group.lift.of }, + refine free_abelian_group.induction_on x rfl _ _ _, + { intros s1, iterate 3 { rw free_abelian_group.lift.of }, rw add_assoc }, + { intros s1 ih, iterate 3 { rw free_abelian_group.lift.neg }, rw ih }, + { intros x1 x2 ih1 ih2, iterate 3 { rw free_abelian_group.lift.add }, rw [ih1, ih2] } }, + { intros s2 ih, iterate 4 { rw free_abelian_group.lift.neg }, rw ih }, + { intros y1 y2 ih1 ih2, iterate 4 { rw free_abelian_group.lift.add }, rw [ih1, ih2] } }, + { intros s3 ih, iterate 3 { rw free_abelian_group.lift.neg }, rw ih }, + { intros z1 z2 ih1 ih2, iterate 2 { rw free_abelian_group.lift.add }, rw [ih1, ih2], + exact (free_abelian_group.lift.add _ _ _).symm } + end } + +instance : ring (free_comm_ring α) := +{ one := free_abelian_group.of 0, + mul_one := λ x, begin + unfold has_mul.mul semigroup.mul has_one.one, + rw free_abelian_group.lift.of, + refine free_abelian_group.induction_on x rfl _ _ _, + { intros s, rw [free_abelian_group.lift.of, add_zero] }, + { intros s ih, rw [free_abelian_group.lift.neg, ih] }, + { intros x1 x2 ih1 ih2, rw [free_abelian_group.lift.add, ih1, ih2] } + end, + one_mul := λ x, begin + unfold has_mul.mul semigroup.mul has_one.one, + refine free_abelian_group.induction_on x rfl _ _ _, + { intros s, rw [free_abelian_group.lift.of, free_abelian_group.lift.of, zero_add] }, + { intros s ih, rw [free_abelian_group.lift.neg, ih] }, + { intros x1 x2 ih1 ih2, rw [free_abelian_group.lift.add, ih1, ih2] } + end, + left_distrib := λ x y z, free_abelian_group.lift.add _ _ _, + right_distrib := λ x y z, begin + unfold has_mul.mul semigroup.mul, + refine free_abelian_group.induction_on z rfl _ _ _, + { intros s, iterate 3 { rw free_abelian_group.lift.of }, rw free_abelian_group.lift.add, refl }, + { intros s ih, iterate 3 { rw free_abelian_group.lift.neg }, rw [ih, neg_add], refl }, + { intros z1 z2 ih1 ih2, iterate 3 { rw free_abelian_group.lift.add }, rw [ih1, ih2], + rw [add_assoc, add_assoc], congr' 1, apply add_left_comm } + end, + .. free_comm_ring.add_comm_group α, + .. free_comm_ring.semigroup α } + +instance : comm_ring (free_comm_ring α) := +{ mul_comm := λ x y, begin + refine free_abelian_group.induction_on x (zero_mul y) _ _ _, + { intros s, refine free_abelian_group.induction_on y (zero_mul _).symm _ _ _, + { intros t, unfold has_mul.mul semigroup.mul ring.mul, + iterate 4 { rw free_abelian_group.lift.of }, rw add_comm }, + { intros t ih, rw [mul_neg_eq_neg_mul_symm, ih, neg_mul_eq_neg_mul] }, + { intros y1 y2 ih1 ih2, rw [mul_add, add_mul, ih1, ih2] } }, + { intros s ih, rw [neg_mul_eq_neg_mul_symm, ih, neg_mul_eq_mul_neg] }, + { intros x1 x2 ih1 ih2, rw [add_mul, mul_add, ih1, ih2] } + end + .. free_comm_ring.ring α } + +variables {α} +def of (x : α) : free_comm_ring α := +free_abelian_group.of [x] + +section lift + +variables {β : Type v} [comm_ring β] (f : α → β) + +def lift : free_comm_ring α → β := +free_abelian_group.lift $ λ s, (s.map f).prod + +@[simp] lemma lift_zero : lift f 0 = 0 := rfl + +@[simp] lemma lift_one : lift f 1 = 1 := +free_abelian_group.lift.of _ _ + +@[simp] lemma lift_of (x : α) : lift f (of x) = f x := +(free_abelian_group.lift.of _ _).trans $ mul_one _ + +@[simp] lemma lift_add (x y) : lift f (x + y) = lift f x + lift f y := +free_abelian_group.lift.add _ _ _ + +@[simp] lemma lift_neg (x) : lift f (-x) = -lift f x := +free_abelian_group.lift.neg _ _ + +@[simp] lemma lift_sub (x y) : lift f (x - y) = lift f x - lift f y := +free_abelian_group.lift.sub _ _ _ + +@[simp] lemma lift_mul (x y) : lift f (x * y) = lift f x * lift f y := +begin + refine free_abelian_group.induction_on y (mul_zero _).symm _ _ _, + { intros s2, conv { to_lhs, dsimp only [(*), mul_zero_class.mul, semiring.mul, ring.mul, semigroup.mul] }, + rw [free_abelian_group.lift.of, lift, free_abelian_group.lift.of], + refine free_abelian_group.induction_on x (zero_mul _).symm _ _ _, + { intros s1, iterate 3 { rw free_abelian_group.lift.of }, rw [multiset.map_add, multiset.prod_add] }, + { intros s1 ih, iterate 3 { rw free_abelian_group.lift.neg }, rw [ih, neg_mul_eq_neg_mul] }, + { intros x1 x2 ih1 ih2, iterate 3 { rw free_abelian_group.lift.add }, rw [ih1, ih2, add_mul] } }, + { intros s2 ih, rw [mul_neg_eq_neg_mul_symm, lift_neg, lift_neg, mul_neg_eq_neg_mul_symm, ih] }, + { intros y1 y2 ih1 ih2, rw [mul_add, lift_add, lift_add, mul_add, ih1, ih2] }, +end + +instance : is_ring_hom (lift f) := +{ map_one := lift_one f, + map_mul := lift_mul f, + map_add := lift_add f } + +@[simp] lemma lift_pow (x) (n : ℕ) : lift f (x ^ n) = lift f x ^ n := +is_semiring_hom.map_pow _ x n + +theorem lift_unique (f : free_comm_ring α → β) [is_ring_hom f] : f = lift (f ∘ of) := +funext $ λ x, @@free_abelian_group.lift.ext _ _ _ + (is_ring_hom.is_add_group_hom f) + (is_ring_hom.is_add_group_hom $ lift $ f ∘ of) + (λ s, multiset.induction_on s ((is_ring_hom.map_one f).trans $ eq.symm $ lift_one _) + (λ hd tl ih, show f (of hd * free_abelian_group.of tl) = lift (f ∘ of) (of hd * free_abelian_group.of tl), + by rw [is_ring_hom.map_mul f, lift_mul, lift_of, ih])) + +end lift + +variables {β : Type v} (f : α → β) + +def map : free_comm_ring α → free_comm_ring β := +lift $ of ∘ f + +instance : monad free_comm_ring := +{ map := λ _ _, map, + pure := λ _, of, + bind := λ _ _ x f, lift f x } + +end free_comm_ring From 8fbae04db1131643b2cedcab1146fe11a4213d81 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Fri, 22 Feb 2019 03:15:50 +0000 Subject: [PATCH 11/22] direct limit of fields (WIP) --- src/algebra/direct_limit.lean | 51 +++++- src/group_theory/free_abelian_group.lean | 32 ++++ src/ring_theory/to_field.lean | 192 +++++++++++++++++++++++ 3 files changed, 272 insertions(+), 3 deletions(-) create mode 100644 src/ring_theory/to_field.lean diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index d9ec27f44b7bd..cc7acc65301d1 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -2,6 +2,7 @@ import linear_algebra.direct_sum_module import linear_algebra.linear_combination import ring_theory.free_comm_ring import ring_theory.ideal_operations +import ring_theory.to_field universes u v w u₁ @@ -11,7 +12,7 @@ class directed_order (α : Type u) extends partial_order α := (directed : ∀ i j : α, ∃ k, i ≤ k ∧ j ≤ k) variables {R : Type u} [ring R] -variables {ι : Type v} [inhabited ι] +variables {ι : Type v} [nonempty ι] variables [directed_order ι] [decidable_eq ι] variables (G : ι → Type w) [Π i, decidable_eq (G i)] @@ -79,8 +80,9 @@ quotient.induction_on' x $ λ y, direct_sum.induction_on y ihx.symm ▸ ihy.symm ▸ (linear_map.map_add _ _ _).symm) theorem exists_of (z : direct_limit G f) : ∃ i x, z = of R ι G f i x := +nonempty.elim (by apply_instance) $ assume ind : ι, quotient.induction_on' z $ λ z, direct_sum.induction_on z - ⟨default _, 0, (linear_map.map_zero _).symm⟩ + ⟨ind, 0, (linear_map.map_zero _).symm⟩ (λ i x, ⟨i, x, rfl⟩) (λ p q ⟨i, x, ihx⟩ ⟨j, y, ihy⟩, let ⟨k, hik, hjk⟩ := directed_order.directed i j in ⟨k, f i k hik x + f j k hjk y, by rw [linear_map.map_add, of_f, of_f, ← ihx, ← ihy]; refl⟩) @@ -124,6 +126,7 @@ end lemma of.zero_exact_aux {x : direct_sum ι G} (H : x ∈ span R (thing G f)) : ∃ j, (∀ k ∈ x.support, k ≤ j) ∧ direct_sum.to_module R ι (G j) (λ i, totalize G f i j) x = (0 : G j) := +nonempty.elim (by apply_instance) $ assume ind : ι, span_induction H (λ x hx, let ⟨i, j, hij, y, hxy⟩ := mem_thing.1 hx in let ⟨k, hik, hjk⟩ := directed_order.directed i j in @@ -140,7 +143,7 @@ span_induction H directed_system.Hcomp f, direct_sum.apply_eq_component, direct_sum.component.of], end⟩) - ⟨default ι, λ _ h, (finset.not_mem_empty _ h).elim, linear_map.map_zero _⟩ + ⟨ind, λ _ h, (finset.not_mem_empty _ h).elim, linear_map.map_zero _⟩ (λ x y ⟨i, hi, hxi⟩ ⟨j, hj, hyj⟩, let ⟨k, hik, hjk⟩ := directed_order.directed i j in ⟨k, λ l hl, @@ -289,6 +292,26 @@ instance of.is_ring_hom (i) : is_ring_hom (of G f i) := map_add := λ x y, ideal.quotient.eq.2 $ subset_span $ or.inl $ or.inr $ set.mem_Union.2 ⟨i, set.mem_Union.2 ⟨x, set.mem_range_self y⟩⟩ } +theorem of_f {i j} (hij : i ≤ j) (x : G i) : of G f j (f i j hij x) = of G f i x := +ideal.quotient.eq.2 $ subset_span $ or.inl $ or.inl $ or.inl $ set.mem_Union.2 ⟨i, set.mem_Union.2 +⟨j, set.mem_Union.2 ⟨hij, set.mem_range_self _⟩⟩⟩ + +theorem exists_of (z : direct_limit G f) : ∃ i x, of G f i x = z := +nonempty.elim (by apply_instance) $ assume ind : ι, +quotient.induction_on' z $ λ x, free_abelian_group.induction_on x + ⟨ind, 0, is_ring_hom.map_zero (of G f ind)⟩ + (λ s, multiset.induction_on s + ⟨ind, 1, is_ring_hom.map_one (of G f ind)⟩ + (λ a s ih, let ⟨i, x⟩ := a, ⟨j, y, hs⟩ := ih, ⟨k, hik, hjk⟩ := directed_order.directed i j in + ⟨k, f i k hik x * f j k hjk y, by rw [is_ring_hom.map_mul (of G f k), of_f, of_f, hs]; refl⟩)) + (λ s ⟨i, x, ih⟩, ⟨i, -x, by rw [is_ring_hom.map_neg (of G f i), ih]; refl⟩) + (λ p q ⟨i, x, ihx⟩ ⟨j, y, ihy⟩, let ⟨k, hik, hjk⟩ := directed_order.directed i j in + ⟨k, f i k hik x + f j k hjk y, by rw [is_ring_hom.map_add (of G f k), of_f, of_f, ihx, ihy]; refl⟩) + +@[elab_as_eliminator] theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit G f) + (ih : ∀ i x, C (of G f i x)) : C z := +let ⟨i, x, hx⟩ := exists_of G f z in hx ▸ ih i x + variables (P : Type u₁) [comm_ring P] variables (g : Π i, G i → P) [Π i, is_ring_hom (g i)] variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) @@ -319,3 +342,25 @@ instance rec.is_ring_hom : is_ring_hom (rec G f P g Hg) := end direct_limit end ring + + +namespace field + +variables [Π i, field (G i)] +variables (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_field_hom (f i j hij)] +variables [add_comm_group.directed_system G f] + +instance direct_limit.nonzero_comm_ring : nonzero_comm_ring (ring.direct_limit G f) := +{ zero_ne_one := sorry, + .. ring.direct_limit.comm_ring G f } + +instance direct_limit.is_division_ring : is_division_ring (ring.direct_limit G f) := +{ exists_inv := λ x, sorry } + +def direct_limit : Type (max v w) := +to_field (ring.direct_limit G f) + +instance : field (direct_limit G f) := +@to_field.field (ring.direct_limit G f) (field.direct_limit.nonzero_comm_ring G f) (field.direct_limit.is_division_ring G f) + +end field diff --git a/src/group_theory/free_abelian_group.lean b/src/group_theory/free_abelian_group.lean index bf589c1642475..f620a72f2b0de 100644 --- a/src/group_theory/free_abelian_group.lean +++ b/src/group_theory/free_abelian_group.lean @@ -115,4 +115,36 @@ begin ac_refl } end +section count + +variables [decidable_eq α] + +def count (x : α) : free_abelian_group α → ℤ := +lift $ λ y, if x = y then 1 else 0 + +theorem count_of (x y : α) : count x (of y) = if x = y then 1 else 0 := +lift.of _ _ + +theorem count_of_self (x : α) : count x (of x) = 1 := +(count_of x x).trans $ if_pos rfl + +theorem count_zero (x : α) : count x 0 = 0 := +lift.zero _ + +end count + +theorem of_inj (x y : α) (H : of x = of y) : x = y := +begin + classical, replace H := congr_arg (count x) H, + rw [count_of_self x, count_of] at H, split_ifs at H with h, + { exact h }, { exfalso, exact one_ne_zero H } +end + +theorem of_ne_zero (x : α) : of x ≠ 0 := +begin + intros H, classical, replace H := congr_arg (count x) H, + rw [count_of_self x, count_zero] at H, + exact one_ne_zero H +end + end free_abelian_group diff --git a/src/ring_theory/to_field.lean b/src/ring_theory/to_field.lean new file mode 100644 index 0000000000000..2cf6f60aa8aae --- /dev/null +++ b/src/ring_theory/to_field.lean @@ -0,0 +1,192 @@ +import algebra.ring + +universes u + +variables (α : Type u) + +class is_division_ring [ring α] : Prop := +(exists_inv : ∀ {x : α}, x ≠ 0 → ∃ y, x * y = 1) + +theorem mul_eq_one_comm {α} [ring α] [is_division_ring α] {x y : α} (hxy : x * y = 1) : y * x = 1 := +classical.by_cases + (assume hy0 : y = 0, by rw [← hxy, hy0, zero_mul, mul_zero]) + (assume hy0 : y ≠ 0, let ⟨z, hz⟩ := is_division_ring.exists_inv hy0 in + by rw [← mul_one x, ← hz, ← mul_assoc x, hxy, one_mul]) + +inductive pre_to_field : Type u +| of : α → pre_to_field +| add : pre_to_field → pre_to_field → pre_to_field +| mul : pre_to_field → pre_to_field → pre_to_field +| inv : pre_to_field → pre_to_field + +namespace pre_to_field + +variables {α} + +def rel [semiring α] : (pre_to_field α) → α → Prop +| (of p) z := p = z +| (add p q) z := ∃ x y, rel p x ∧ rel q y ∧ x + y = z +| (mul p q) z := ∃ x y, rel p x ∧ rel q y ∧ x * y = z +| (inv p) z := rel p 0 ∧ z = 0 ∨ ∃ x, rel p x ∧ x * z = 1 + +def neg [has_neg α] : pre_to_field α → pre_to_field α +| (of p) := of (-p) +| (add x y) := add (neg x) (neg y) +| (mul x y) := mul (neg x) y +| (inv x) := inv (neg x) + +theorem rel_neg [ring α] : Π {p} {z : α}, rel p z → rel (neg p) (-z) +| (of p) z hpz := congr_arg has_neg.neg hpz +| (add p q) z ⟨x, y, hpx, hqy, hxyz⟩ := ⟨-x, -y, rel_neg hpx, rel_neg hqy, by rw [← hxyz, neg_add]⟩ +| (mul p q) z ⟨x, y, hpx, hqy, hxyz⟩ := ⟨-x, y, rel_neg hpx, hqy, by rw [← hxyz, neg_mul_eq_neg_mul]⟩ +| (inv p) z (or.inl ⟨hp0, hz0⟩) := or.inl ⟨@neg_zero α _ ▸ rel_neg hp0, neg_eq_zero.2 hz0⟩ +| (inv p) z (or.inr ⟨x, hpx, hxz⟩) := or.inr ⟨-x, rel_neg hpx, by rw [neg_mul_neg, hxz]⟩ + +variables [ring α] [is_division_ring α] + +lemma unique_rel (p : pre_to_field α) : + ∃!x, rel p x := +pre_to_field.rec_on p + (λ x, ⟨x, rfl, λ y hxy, hxy.symm⟩) + (λ p q ⟨x, hpx, hpx2⟩ ⟨y, hqy, hqy2⟩, ⟨x+y, ⟨x, y, hpx, hqy, rfl⟩, + λ z ⟨x', y', hpx', hqy', hxyz⟩, hpx2 x' hpx' ▸ hqy2 y' hqy' ▸ hxyz.symm⟩) + (λ p q ⟨x, hpx, hpx2⟩ ⟨y, hqy, hqy2⟩, ⟨x*y, ⟨x, y, hpx, hqy, rfl⟩, + λ z ⟨x', y', hpx', hqy', hxyz⟩, hpx2 x' hpx' ▸ hqy2 y' hqy' ▸ hxyz.symm⟩) + (λ p ⟨x, hpx, hpx2⟩, classical.by_cases + (assume hx0 : x = 0, ⟨0, or.inl ⟨hx0 ▸ hpx, rfl⟩, λ y hy, + or.cases_on hy and.right $ λ ⟨z, hpz, hzy⟩, + by rw [← one_mul y, ← hzy, hpx2 z hpz, hx0, zero_mul, zero_mul]⟩) + (assume hx0 : x ≠ 0, let ⟨y, hxy⟩ := is_division_ring.exists_inv hx0 in + ⟨y, or.inr ⟨x, hpx, hxy⟩, λ z hz, or.cases_on hz + (λ ⟨hp0, hz0⟩, by rw [hz0, ← one_mul y, ← hxy, ← hpx2 0 hp0, zero_mul, zero_mul]) + (λ ⟨s, hps, hsz⟩, by rw [← mul_one y, ← hsz, ← mul_assoc, hpx2 s hps, mul_eq_one_comm hxy, one_mul])⟩)) + +variables (α) + +instance : setoid (pre_to_field α) := +{ r := λ p q, ∃ x, rel p x ∧ rel q x, + iseqv := ⟨λ p, let ⟨x, hx⟩ := unique_rel p in ⟨x, hx.1, hx.1⟩, + λ p q ⟨x, hpx, hqx⟩, ⟨x, hqx, hpx⟩, + λ p q r ⟨x, hpx, hqx⟩ ⟨y, hqy, hry⟩, ⟨x, hpx, unique_of_exists_unique (unique_rel q) hqy hqx ▸ hry⟩⟩ } + +end pre_to_field + + +def to_field [ring α] [is_division_ring α] : Type u := +quotient (pre_to_field.setoid α) + +namespace to_field + +variables {α} [ring α] [is_division_ring α] + +def mk (x : α) : to_field α := +⟦pre_to_field.of x⟧ + +variables (α) + +theorem mk.bijective : function.bijective (@mk α _ _) := +⟨λ x y H, let ⟨z, hxz, hyz⟩ := quotient.exact H in eq.trans hxz hyz.symm, +λ x, quotient.induction_on x $ λ p, let ⟨z, hz⟩ := pre_to_field.unique_rel p in +⟨z, quotient.sound ⟨z, rfl, hz.1⟩⟩⟩ + +instance : has_add (to_field α) := +⟨λ x y, quotient.lift_on₂ x y (λ p q, ⟦pre_to_field.add p q⟧) $ λ p q r s ⟨x, hpx, hrx⟩ ⟨y, hqy, hsy⟩, quotient.sound +⟨x + y, ⟨x, y, hpx, hqy, rfl⟩, ⟨x, y, hrx, hsy, rfl⟩⟩⟩ + +instance : has_neg (to_field α) := +⟨λ x, quotient.lift_on x (λ p, ⟦pre_to_field.neg p⟧) $ λ p q ⟨x, hpx, hqx⟩, quotient.sound +⟨-x, pre_to_field.rel_neg hpx, pre_to_field.rel_neg hqx⟩⟩ + +instance : has_mul (to_field α) := +⟨λ x y, quotient.lift_on₂ x y (λ p q, ⟦pre_to_field.mul p q⟧) $ λ p q r s ⟨x, hpx, hrx⟩ ⟨y, hqy, hsy⟩, quotient.sound +⟨x * y, ⟨x, y, hpx, hqy, rfl⟩, ⟨x, y, hrx, hsy, rfl⟩⟩⟩ + +instance : has_inv (to_field α) := +⟨λ x, quotient.lift_on x (λ p, ⟦pre_to_field.inv p⟧) $ λ p q ⟨x, hpx, hqx⟩, quotient.sound $ +classical.by_cases + (assume hx0 : x = 0, ⟨0, or.inl ⟨hx0 ▸ hpx, rfl⟩, or.inl ⟨hx0 ▸ hqx, rfl⟩⟩) + (assume hx0 : x ≠ 0, let ⟨y, hy⟩ := is_division_ring.exists_inv hx0 in ⟨y, or.inr ⟨x, hpx, hy⟩, or.inr ⟨x, hqx, hy⟩⟩)⟩ + +variables {α} + +@[elab_as_eliminator] protected lemma induction_on + {C : to_field α → Prop} (x : to_field α) + (ih : ∀ z, C (mk z)) : C x := +let ⟨z, hz⟩ := (mk.bijective α).2 x in hz ▸ ih z + +@[elab_as_eliminator] protected lemma induction_on₂ + {C : to_field α → to_field α → Prop} (x y : to_field α) + (ih : ∀ p q, C (mk p) (mk q)) : C x y := +to_field.induction_on x $ λ p, to_field.induction_on y $ λ q, ih p q + +@[elab_as_eliminator] protected lemma induction_on₃ + {C : to_field α → to_field α → to_field α → Prop} (x y z : to_field α) + (ih : ∀ p q r, C (mk p) (mk q) (mk r)) : C x y z := +to_field.induction_on x $ λ p, to_field.induction_on y $ λ q, to_field.induction_on z $ λ r, ih p q r + +@[simp] lemma mk_add (x y : α) : mk (x + y) = mk x + mk y := +quotient.sound ⟨x + y, rfl, ⟨x, y, rfl, rfl, rfl⟩⟩ + +@[simp] lemma mk_neg (x : α) : mk (-x) = -mk x := rfl + +@[simp] lemma mk_mul (x y : α) : mk (x * y) = mk x * mk y := +quotient.sound ⟨x * y, rfl, ⟨x, y, rfl, rfl, rfl⟩⟩ + +lemma mk_inv {x y : α} (hxy : x * y = 1) : (mk x)⁻¹ = mk y := +quotient.sound ⟨y, or.inr ⟨x, rfl, hxy⟩, rfl⟩ + +variables (α) +instance : ring (to_field α) := +{ zero := mk 0, + one := mk 1, + add_assoc := λ x y z, to_field.induction_on₃ x y z $ λ p q r, + by rw [← mk_add, ← mk_add, ← mk_add, ← mk_add, add_assoc], + zero_add := λ x, to_field.induction_on x $ λ p, + by rw [← mk_add, zero_add], + add_zero := λ x, to_field.induction_on x $ λ p, + by rw [← mk_add, add_zero], + add_left_neg := λ x, to_field.induction_on x $ λ p, + by rw [← mk_neg, ← mk_add, add_left_neg]; refl, + add_comm := λ x y, to_field.induction_on₂ x y $ λ p q, + by rw [← mk_add, ← mk_add, add_comm], + mul_assoc := λ x y z, to_field.induction_on₃ x y z $ λ p q r, + by rw [← mk_mul, ← mk_mul, ← mk_mul, ← mk_mul, mul_assoc], + one_mul := λ x, to_field.induction_on x $ λ p, + by rw [← mk_mul, one_mul], + mul_one := λ x, to_field.induction_on x $ λ p, + by rw [← mk_mul, mul_one], + left_distrib := λ x y z, to_field.induction_on₃ x y z $ λ p q r, + by rw [← mk_add, ← mk_mul, ← mk_mul, ← mk_mul, ← mk_add, mul_add], + right_distrib := λ x y z, to_field.induction_on₃ x y z $ λ p q r, + by rw [← mk_add, ← mk_mul, ← mk_mul, ← mk_mul, ← mk_add, add_mul], + .. to_field.has_add α, + .. to_field.has_neg α, + .. to_field.has_mul α } + +@[simp] lemma mk_zero : mk (0 : α) = 0 := rfl +@[simp] lemma mk_one : mk (1 : α) = 1 := rfl + +instance : is_ring_hom (@mk α _ _) := +⟨mk_one α, mk_mul, mk_add⟩ + +end to_field + +namespace to_field + +instance [comm_ring α] [is_division_ring α] : comm_ring (to_field α) := +{ mul_comm := λ x y, to_field.induction_on₂ x y $ λ p q, + by rw [← mk_mul, ← mk_mul, mul_comm], + .. to_field.ring α } + +instance [nonzero_comm_ring α] [is_division_ring α] : field (to_field α) := +{ zero_ne_one := λ H, zero_ne_one $ (mk.bijective α).1 H, + mul_inv_cancel := λ x, to_field.induction_on x $ λ p hp0, + let ⟨y, hy⟩ := is_division_ring.exists_inv (mt (congr_arg mk) hp0) in + by rw [mk_inv hy, ← mk_mul, hy, mk_one], + inv_mul_cancel := λ x, to_field.induction_on x $ λ p hp0, + let ⟨y, hy⟩ := is_division_ring.exists_inv (mt (congr_arg mk) hp0) in + by rw [mk_inv hy, ← mk_mul, mul_eq_one_comm hy, mk_one], + .. to_field.comm_ring α, + .. to_field.has_inv α } + +end to_field From 0617fd9fa8becbfe7d3cb96ecb09c861c421fbeb Mon Sep 17 00:00:00 2001 From: kckennylau Date: Sat, 23 Feb 2019 00:40:51 +0000 Subject: [PATCH 12/22] trying to prove zero_exact for rings --- src/algebra/direct_limit.lean | 152 +++++++++++++++++++- src/ring_theory/free_comm_ring.lean | 38 ++++- src/ring_theory/free_comm_ring_support.lean | 146 +++++++++++++++++++ src/ring_theory/to_field.lean | 43 +++++- 4 files changed, 370 insertions(+), 9 deletions(-) create mode 100644 src/ring_theory/free_comm_ring_support.lean diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index cc7acc65301d1..7d2139afce0e1 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -1,6 +1,6 @@ import linear_algebra.direct_sum_module import linear_algebra.linear_combination -import ring_theory.free_comm_ring +import ring_theory.free_comm_ring_support import ring_theory.ideal_operations import ring_theory.to_field @@ -21,7 +21,7 @@ namespace module variables [Π i, add_comm_group (G i)] [Π i, module R (G i)] class directed_system (f : Π i j, i ≤ j → G i →ₗ[R] G j) : Prop := -(Hid : ∀ i x, f i i (le_refl i) x = x) +(Hid : ∀ i x h, f i i h x = x) (Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) variables (f : Π i j, i ≤ j → G i →ₗ[R] G j) [directed_system G f] @@ -218,7 +218,7 @@ namespace add_comm_group variables [Π i, add_comm_group (G i)] class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := -(Hid : ∀ i x, f i i (le_refl i) x = x) +(Hid : ∀ i x h, f i i h x = x) (Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) def direct_limit @@ -282,6 +282,9 @@ namespace direct_limit instance : comm_ring (direct_limit G f) := ideal.quotient.comm_ring _ +instance : ring (direct_limit G f) := +comm_ring.to_ring _ + def of (i) (x : G i) : direct_limit G f := ideal.quotient.mk _ $ of ⟨i, x⟩ @@ -334,6 +337,8 @@ ideal.quotient.lift _ (free_comm_ring.lift $ λ x, g x.1 x.2) begin { rcases hx with ⟨i, x, y, rfl⟩, simp only [lift_sub, lift_of, lift_mul, is_ring_hom.map_mul (g i), sub_self] } end +omit Hg + instance rec.is_ring_hom : is_ring_hom (rec G f P g Hg) := ⟨free_comm_ring.lift_one _, λ x y, quotient.induction_on₂' x y $ λ p q, free_comm_ring.lift_mul _ _ _, @@ -344,6 +349,118 @@ end direct_limit end ring +namespace ring + +namespace direct_limit + +variables [Π i, comm_ring (G i)] +variables (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_ring_hom (f i j hij)] +variables [add_comm_group.directed_system G f] + +open free_comm_ring + +/-section totalize +local attribute [instance, priority 0] classical.dec + +noncomputable def totalize : Π i j, G i → G j := +λ i j, if h : i ≤ j then f i j h else 0 + +instance totalize.is_ring_hom (i j) : is_ring_hom (totalize G f i j) := +by unfold totalize; split_ifs; apply_instance + +lemma totalize_apply (i j x) : + totalize G f i j x = if h : i ≤ j then f i j h x else 0 := +if h : i ≤ j + then by dsimp only [totalize]; rw [dif_pos h, dif_pos h] + else by dsimp only [totalize]; rw [dif_neg h, dif_neg h, pi.zero_apply] +end totalize-/ + +/-lemma to_module_totalize_of_le {x : free_comm_ring Σ i, G i} {i j : ι} + (hij : i ≤ j) (hx : ∀ k ∈ x.support, k ≤ i) : + direct_sum.to_module R ι (G j) (λ k, totalize G f k j) x = + f i j hij (direct_sum.to_module R ι (G i) (λ k, totalize G f k i) x) := +begin + rw [← @dfinsupp.sum_single ι G _ _ _ x, dfinsupp.sum], + simp only [linear_map.map_sum], + refine finset.sum_congr rfl (λ k hk, _), + rw direct_sum.single_eq_lof R k (x k), + simp [totalize_apply, hx k hk, le_trans (hx k hk) hij, directed_system.Hcomp f] +end-/ + +section of_zero_exact_aux +attribute [instance, priority 0] classical.dec +lemma of.zero_exact_aux {x : free_comm_ring Σ i, G i} (H : x ∈ ideal.span (thing1 G f ∪ thing2 G ∪ thing3 G ∪ thing4 G)) : + ∃ j s, ∃ H : (∀ k : Σ i, G i, k ∈ s → k.1 ≤ j), is_supported x s ∧ + lift (λ ix : s, f ix.1.1 j (H ix ix.2) ix.1.2) (restriction s x) = (0 : G j) := +begin + refine span_induction H _ _ _ _, + { rintros x (⟨⟨hx | hx⟩ | hx⟩ | hx), + { simp only [thing1, set.mem_Union, set.mem_range] at hx, rcases hx with ⟨i, j, hij, x, rfl⟩, + refine ⟨j, {⟨i, x⟩, ⟨j, f i j hij x⟩}, _, + is_supported_sub (is_supported_pure.2 $ or.inl rfl) (is_supported_pure.2 $ or.inr $ or.inl rfl), _⟩, + { rintros k (rfl | ⟨rfl | h⟩), refl, exact hij, cases h }, + { rw [restriction_sub, lift_sub, restriction_of, dif_pos, restriction_of, dif_pos, lift_pure, lift_pure], + dsimp only, rw add_comm_group.directed_system.Hcomp f, exact sub_self _, + { left, refl }, { right, left, refl }, } }, + { simp only [thing2, set.mem_range] at hx, rcases hx with ⟨i, rfl⟩, + refine ⟨i, {⟨i, 1⟩}, _, is_supported_sub (is_supported_pure.2 $ or.inl rfl) is_supported_one, _⟩, + { rintros k (rfl | h), refl, cases h }, + { rw [restriction_sub, lift_sub, restriction_of, dif_pos, restriction_one, lift_pure, lift_one], + dsimp only, rw [is_ring_hom.map_one (f i i _), sub_self], exact _inst_7 i i _, { left, refl } } }, + { simp only [thing3, set.mem_Union, set.mem_range] at hx, rcases hx with ⟨i, x, y, rfl⟩, + refine ⟨i, {⟨i, x+y⟩, ⟨i, x⟩, ⟨i, y⟩}, _, + is_supported_sub (is_supported_pure.2 $ or.inr $ or.inr $ or.inl rfl) + (is_supported_add (is_supported_pure.2 $ or.inr $ or.inl rfl) (is_supported_pure.2 $ or.inl rfl)), _⟩, + { rintros k (rfl | ⟨rfl | ⟨rfl | hk⟩⟩), refl, refl, refl, cases hk }, + { rw [restriction_sub, restriction_add, restriction_of, restriction_of, restriction_of, + dif_pos, dif_pos, dif_pos, lift_sub, lift_add, lift_pure, lift_pure, lift_pure], + dsimp only, simp only [add_comm_group.directed_system.Hid f], rw sub_self, + { left, refl }, { right, left, refl }, { right, right, left, refl } } }, + { simp only [thing4, set.mem_Union, set.mem_range] at hx, rcases hx with ⟨i, x, y, rfl⟩, + refine ⟨i, {⟨i, x*y⟩, ⟨i, x⟩, ⟨i, y⟩}, _, + is_supported_sub (is_supported_pure.2 $ or.inr $ or.inr $ or.inl rfl) + (is_supported_mul (is_supported_pure.2 $ or.inr $ or.inl rfl) (is_supported_pure.2 $ or.inl rfl)), _⟩, + { rintros k (rfl | ⟨rfl | ⟨rfl | hk⟩⟩), refl, refl, refl, cases hk }, + { rw [restriction_sub, restriction_mul, restriction_of, restriction_of, restriction_of, + dif_pos, dif_pos, dif_pos, lift_sub, lift_mul, lift_pure, lift_pure, lift_pure], + dsimp only, simp only [add_comm_group.directed_system.Hid f], rw sub_self, + { left, refl }, { right, left, refl }, { right, right, left, refl } } } }, + { refine nonempty.elim (by apply_instance) (assume ind : ι, _), + refine ⟨ind, ∅, λ _, false.elim, is_supported_zero, _⟩, + rw [restriction_zero, lift_zero] }, + { rintros x y ⟨i, s, hi, hxs, ihs⟩ ⟨j, t, hj, hxt, iht⟩, + rcases directed_order.directed i j with ⟨k, hik, hjk⟩, + refine ⟨k, s ∪ t, _, is_supported_add (is_supported_upwards hxs $ set.subset_union_left s t) + (is_supported_upwards hxt $ set.subset_union_right s t), _⟩, + { rintros t (ht | ht), exact le_trans (hi _ ht) hik, exact le_trans (hj _ ht) hjk }, + { rw [restriction_add, lift_add], sorry } }, + sorry +end +end of_zero_exact_aux + +lemma of.zero_exact {i x} (hix : of G f i x = 0) : ∃ j, ∃ hij : i ≤ j, f i j hij x = 0 := +let ⟨j, s, H, hxs, hx⟩ := of.zero_exact_aux G f (ideal.quotient.eq_zero_iff_mem.1 hix) in +have hixs : (⟨i, x⟩ : Σ i, G i) ∈ s, from is_supported_pure.1 hxs, +⟨j, H ⟨i, x⟩ hixs, by rw [restriction_of, dif_pos hixs, lift_pure] at hx; exact hx⟩ +/- +theorem of_inj (hf : ∀ i j hij, function.injective (f i j hij)) (i) : + function.injective (of G f i) := +begin + suffices : ∀ x, of G f i x = 0 → x = 0, + { intros x y hxy, rw ← sub_eq_zero_iff_eq, apply this, + rw [is_ring_hom.map_sub (of G f i), hxy, sub_self] }, + intros x hx, replace hx := ideal.quotient.eq_zero_iff_mem.1 hx, + generalize_hyp hy : free_comm_ring.of _ = y at hx, revert hy, + refine span_induction hx _ _ _ _, + { rintros z ⟨⟨hz | hz⟩ | hz⟩, + { simp only [thing1, set.mem_Union, set.mem_range] at hz } } +end-/ + +end direct_limit + +end ring + + namespace field variables [Π i, field (G i)] @@ -355,12 +472,39 @@ instance direct_limit.nonzero_comm_ring : nonzero_comm_ring (ring.direct_limit G .. ring.direct_limit.comm_ring G f } instance direct_limit.is_division_ring : is_division_ring (ring.direct_limit G f) := -{ exists_inv := λ x, sorry } +{ exists_inv := λ p, ring.direct_limit.induction_on G f p $ λ i x H, + ⟨ring.direct_limit.of G f i (x⁻¹), by erw [← is_ring_hom.map_mul (ring.direct_limit.of G f i), + mul_inv_cancel (assume h : x = 0, H $ by rw [h, is_ring_hom.map_zero (ring.direct_limit.of G f i)]), + is_ring_hom.map_one (ring.direct_limit.of G f i)]⟩ } def direct_limit : Type (max v w) := to_field (ring.direct_limit G f) +namespace direct_limit + +set_option class.instance_max_depth 10 instance : field (direct_limit G f) := @to_field.field (ring.direct_limit G f) (field.direct_limit.nonzero_comm_ring G f) (field.direct_limit.is_division_ring G f) +def of (i) (x : G i) : direct_limit G f := +to_field.mk $ ring.direct_limit.of G f i x + +set_option class.instance_max_depth 20 +instance of.is_ring_hom (i) : is_ring_hom (of G f i) := +@is_ring_hom.comp _ _ _ _ (ring.direct_limit.of G f i) (ring.direct_limit.of.is_ring_hom G f i) + _ _ (to_field.mk) (to_field.mk.is_ring_hom _) + +theorem of_f {i j} (hij : i ≤ j) (x : G i) : of G f j (f i j hij x) = of G f i x := +congr_arg to_field.mk $ ring.direct_limit.of_f G f hij x + +variables (P : Type u₁) [field' P] +variables (g : Π i, G i → P) [Π i, is_field_hom (g i)] +variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) +include Hg + +def rec : direct_limit G f → P := +to_field.eval $ ring.direct_limit.rec G f P g Hg + +end direct_limit + end field diff --git a/src/ring_theory/free_comm_ring.lean b/src/ring_theory/free_comm_ring.lean index 20ff8528dbb8d..12775debff2ae 100644 --- a/src/ring_theory/free_comm_ring.lean +++ b/src/ring_theory/free_comm_ring.lean @@ -77,6 +77,19 @@ variables {α} def of (x : α) : free_comm_ring α := free_abelian_group.of [x] +@[elab_as_eliminator] protected lemma induction_on + {C : free_comm_ring α → Prop} (z : free_comm_ring α) + (hn1 : C (-1)) (hb : ∀ b, C (of b)) + (ha : ∀ x y, C x → C y → C (x + y)) + (hm : ∀ x y, C x → C y → C (x * y)) : C z := +have hn : ∀ x, C x → C (-x), from λ x ih, neg_one_mul x ▸ hm _ _ hn1 ih, +have h1 : C 1, from neg_neg (1 : free_comm_ring α) ▸ hn _ hn1, +free_abelian_group.induction_on z + (add_left_neg (1 : free_comm_ring α) ▸ ha _ _ hn1 h1) + (λ m, multiset.induction_on m h1 $ λ a m ih, hm _ _ (hb a) ih) + (λ m ih, hn _ ih) + ha + section lift variables {β : Type v} [comm_ring β] (f : α → β) @@ -132,9 +145,7 @@ funext $ λ x, @@free_abelian_group.lift.ext _ _ _ end lift -variables {β : Type v} (f : α → β) - -def map : free_comm_ring α → free_comm_ring β := +def map {β : Type v} (f : α → β) : free_comm_ring α → free_comm_ring β := lift $ of ∘ f instance : monad free_comm_ring := @@ -142,4 +153,25 @@ instance : monad free_comm_ring := pure := λ _, of, bind := λ _ _ x f, lift f x } +@[simp] lemma lift_pure {β : Type v} [comm_ring β] (f : α → β) (x : α) : lift f (pure x) = f x := +lift_of _ _ + +@[elab_as_eliminator] protected lemma induction_on' + {C : free_comm_ring α → Prop} (z : free_comm_ring α) + (hn1 : C (-1)) (hb : ∀ b, C (pure b)) + (ha : ∀ x y, C x → C y → C (x + y)) + (hm : ∀ x y, C x → C y → C (x * y)) : C z := +free_comm_ring.induction_on z hn1 hb ha hm + +section map +variables {β : Type u} (f : α → β) (x y : free_comm_ring α) +@[simp] lemma map_pure (z) : f <$> (pure z : free_comm_ring α) = pure (f z) := lift_of (of ∘ f) z +@[simp] lemma map_zero : f <$> (0 : free_comm_ring α) = 0 := lift_zero (of ∘ f) +@[simp] lemma map_one : f <$> (1 : free_comm_ring α) = 1 := lift_one (of ∘ f) +@[simp] lemma map_add : f <$> (x + y) = f <$> x + f <$> y := lift_add (of ∘ f) x y +@[simp] lemma map_neg : f <$> (-x) = -(f <$> x) := lift_neg (of ∘ f) x +@[simp] lemma map_sub : f <$> (x - y) = f <$> x - f <$> y := lift_sub (of ∘ f) x y +@[simp] lemma map_mul : f <$> (x * y) = f <$> x * f <$> y := lift_mul (of ∘ f) x y +end map + end free_comm_ring diff --git a/src/ring_theory/free_comm_ring_support.lean b/src/ring_theory/free_comm_ring_support.lean new file mode 100644 index 0000000000000..80762389ffc9c --- /dev/null +++ b/src/ring_theory/free_comm_ring_support.lean @@ -0,0 +1,146 @@ +import ring_theory.free_comm_ring ring_theory.subring data.set.finite data.real.irrational + +@[derive decidable_eq] +inductive F4 : Type | z | o | a | b + +namespace F4 + +instance : fintype F4 := +{ elems := {z, o, a, b}, + complete := λ x, by cases x; simp } + +def add : F4 → F4 → F4 +| z x := x +| x z := x +| o o := z +| o a := b +| o b := a +| a o := b +| a a := z +| a b := o +| b o := a +| b a := o +| b b := z + +def mul : F4 → F4 → F4 +| z x := z +| x z := z +| o x := x +| x o := x +| a a := b +| a b := o +| b a := o +| b b := a + +def inv : F4 → F4 +| z := z +| o := o +| a := b +| b := a + +instance : discrete_field F4 := +by refine +{ add := add, + zero := z, + neg := id, + mul := mul, + one := o, + inv := inv, + has_decidable_eq := F4.decidable_eq, + .. }; tactic.exact_dec_trivial + +end F4 + +universes u v + +variables {α : Type u} + +namespace free_comm_ring + +def is_supported (x : free_comm_ring α) (s : set α) : Prop := +x ∈ ring.closure (of '' s) + +theorem is_supported_upwards {x : free_comm_ring α} {s t} (hs : is_supported x s) (hst : s ⊆ t) : + is_supported x t := +ring.closure_mono (set.mono_image hst) hs + +theorem is_supported_add {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : + is_supported (x + y) s := +is_add_submonoid.add_mem hxs hys + +theorem is_supported_neg {x : free_comm_ring α} {s} (hxs : is_supported x s) : + is_supported (-x) s := +is_add_subgroup.neg_mem hxs + +theorem is_supported_sub {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : + is_supported (x - y) s := +is_add_subgroup.sub_mem _ _ _ hxs hys + +theorem is_supported_mul {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : + is_supported (x * y) s := +is_submonoid.mul_mem hxs hys + +theorem is_supported_zero {s : set α} : is_supported 0 s := +is_add_submonoid.zero_mem _ + +theorem is_supported_one {s : set α} : is_supported 1 s := +is_submonoid.one_mem _ + +theorem is_supported_int {i : ℤ} {s : set α} : is_supported ↑i s := +int.induction_on i is_supported_zero + (λ i hi, by rw [int.cast_add, int.cast_one]; exact is_supported_add hi is_supported_one) + (λ i hi, by rw [int.cast_sub, int.cast_one]; exact is_supported_sub hi is_supported_one) + +def restriction (s : set α) [decidable_pred s] (x : free_comm_ring α) : free_comm_ring s := +lift (λ p, if H : p ∈ s then pure ⟨p, H⟩ else 0) x + +section restriction +variables (s : set α) [decidable_pred s] (x y : free_comm_ring α) +@[simp] lemma restriction_of (p) : restriction s (of p) = if H : p ∈ s then pure ⟨p, H⟩ else 0 := lift_of _ _ +@[simp] lemma restriction_pure (p) : restriction s (pure p) = if H : p ∈ s then pure ⟨p, H⟩ else 0 := lift_of _ _ +@[simp] lemma restriction_zero : restriction s 0 = 0 := lift_zero _ +@[simp] lemma restriction_one : restriction s 1 = 1 := lift_one _ +@[simp] lemma restriction_add : restriction s (x + y) = restriction s x + restriction s y := lift_add _ _ _ +@[simp] lemma restriction_neg : restriction s (-x) = -restriction s x := lift_neg _ _ +@[simp] lemma restriction_sub : restriction s (x - y) = restriction s x - restriction s y := lift_sub _ _ _ +@[simp] lemma restriction_mul : restriction s (x * y) = restriction s x * restriction s y := lift_mul _ _ _ +end restriction + +theorem is_supported_pure {p} {s : set α} : is_supported (pure p) s ↔ p ∈ s := +suffices is_supported (pure p) s → p ∈ s, from ⟨this, λ hps, ring.subset_closure ⟨p, hps, rfl⟩⟩, +assume hps : is_supported (pure p) s, begin + classical, + have : ∀ x, is_supported x s → lift (λ q, if q ∈ s then 0 else F4.a) x = 0 ∨ + lift (λ q, if q ∈ s then 0 else F4.a) x = 1, + { intros x hx, refine ring.in_closure.rec_on hx _ _ _ _, + { right, rw lift_one }, + { right, rw [lift_neg, lift_one], refl }, + { rintros _ ⟨q, hqs, rfl⟩ _ _, left, rw [lift_mul, lift_of, if_pos hqs, zero_mul] }, + { rintros x y (ihx | ihx) (ihy | ihy); rw [lift_add, ihx, ihy], + exacts [or.inl rfl, or.inr rfl, or.inr rfl, or.inl rfl] } }, + specialize this (of p) hps, rw [lift_of] at this, split_ifs at this, { exact h }, + exfalso, cases this; cases this +end + +theorem subtype_val_map_restriction {x} (s : set α) [decidable_pred s] (hxs : is_supported x s) : + subtype.val <$> restriction s x = x := +begin + refine ring.in_closure.rec_on hxs _ _ _ _, + { rw restriction_one, refl }, + { rw [restriction_neg, map_neg, restriction_one], refl }, + { rintros _ ⟨p, hps, rfl⟩ n ih, rw [restriction_mul, restriction_of, dif_pos hps, map_mul, map_pure, ih], refl }, + { intros x y ihx ihy, rw [restriction_add, map_add, ihx, ihy] } +end + +theorem exists_finset (x : free_comm_ring α) : ∃ s : set α, set.finite s ∧ is_supported x s := +free_comm_ring.induction_on x + ⟨∅, set.finite_empty, is_supported_neg is_supported_one⟩ + (λ p, ⟨{p}, set.finite_singleton p, is_supported_pure.2 $ finset.mem_singleton_self _⟩) + (λ x y ⟨s, hfs, hxs⟩ ⟨t, hft, hxt⟩, ⟨s ∪ t, set.finite_union hfs hft, is_supported_add + (is_supported_upwards hxs $ set.subset_union_left s t) + (is_supported_upwards hxt $ set.subset_union_right s t)⟩) + (λ x y ⟨s, hfs, hxs⟩ ⟨t, hft, hxt⟩, ⟨s ∪ t, set.finite_union hfs hft, is_supported_mul + (is_supported_upwards hxs $ set.subset_union_left s t) + (is_supported_upwards hxt $ set.subset_union_right s t)⟩) + +end free_comm_ring diff --git a/src/ring_theory/to_field.lean b/src/ring_theory/to_field.lean index 2cf6f60aa8aae..246729782a7d5 100644 --- a/src/ring_theory/to_field.lean +++ b/src/ring_theory/to_field.lean @@ -1,12 +1,21 @@ import algebra.ring -universes u +universes u v variables (α : Type u) class is_division_ring [ring α] : Prop := (exists_inv : ∀ {x : α}, x ≠ 0 → ∃ y, x * y = 1) +section +set_option old_structure_cmd true +class field' (α : Type u) extends division_ring α, comm_ring α := +(inv_zero : (0:α)⁻¹ = 0) +end + +theorem inv_zero' [field' α] : (0:α)⁻¹ = 0 := +field'.inv_zero α + theorem mul_eq_one_comm {α} [ring α] [is_division_ring α] {x y : α} (hxy : x * y = 1) : y * x = 1 := classical.by_cases (assume hy0 : y = 0, by rw [← hxy, hy0, zero_mul, mul_zero]) @@ -61,6 +70,26 @@ pre_to_field.rec_on p (λ ⟨hp0, hz0⟩, by rw [hz0, ← one_mul y, ← hxy, ← hpx2 0 hp0, zero_mul, zero_mul]) (λ ⟨s, hps, hsz⟩, by rw [← mul_one y, ← hsz, ← mul_assoc, hpx2 s hps, mul_eq_one_comm hxy, one_mul])⟩)) +variables {β : Type v} + +def eval [division_ring β] (f : α → β) : pre_to_field α → β +| (of p) := f p +| (add x y) := eval x + eval y +| (mul x y) := eval x * eval y +| (inv p) := (eval p)⁻¹ + +theorem eval_of_rel [field' β] (f : α → β) [is_ring_hom f] : + Π {p} {x : α} (hpx : pre_to_field.rel p x), eval f p = f x +| (of p) z hpz := congr_arg f hpz +| (add p q) z ⟨x, y, hpx, hqy, hxyz⟩ := by rw [eval, eval_of_rel hpx, eval_of_rel hqy, ← hxyz, is_ring_hom.map_add f] +| (mul p q) z ⟨x, y, hpx, hqy, hxyz⟩ := by rw [eval, eval_of_rel hpx, eval_of_rel hqy, ← hxyz, is_ring_hom.map_mul f] +| (inv p) z (or.inl ⟨hp0, hz0⟩) := by rw [eval, eval_of_rel hp0, hz0, is_ring_hom.map_zero f, inv_zero'] +| (inv p) z (or.inr ⟨x, hpx, hxz⟩) := by rw [eval, eval_of_rel hpx, ← mul_one (f x)⁻¹, ← is_ring_hom.map_one f, + ← hxz, is_ring_hom.map_mul f, ← mul_assoc, + inv_mul_cancel (assume hfx : f x = 0, @zero_ne_one β _ $ + by rw [← is_ring_hom.map_one f, ← hxz, is_ring_hom.map_mul f, hfx, zero_mul]), + one_mul] + variables (α) instance : setoid (pre_to_field α) := @@ -166,9 +195,19 @@ instance : ring (to_field α) := @[simp] lemma mk_zero : mk (0 : α) = 0 := rfl @[simp] lemma mk_one : mk (1 : α) = 1 := rfl -instance : is_ring_hom (@mk α _ _) := +instance mk.is_ring_hom : is_ring_hom (@mk α _ _) := ⟨mk_one α, mk_mul, mk_add⟩ +variables {α} +def eval {β : Type v} [field' β] (f : α → β) [is_ring_hom f] (p : to_field α) : β := +quotient.lift_on p (pre_to_field.eval f) $ λ p q ⟨x, hpx, hqx⟩, +by rw [pre_to_field.eval_of_rel f hpx, pre_to_field.eval_of_rel f hqx] + +instance eval.is_ring_hom {β : Type v} [field' β] (f : α → β) [is_ring_hom f] : is_ring_hom (eval f) := +⟨by convert is_ring_hom.map_one f, +λ p q, quotient.induction_on₂ p q $ λ x y, rfl, +λ p q, quotient.induction_on₂ p q $ λ x y, rfl⟩ + end to_field namespace to_field From 5b578eee16f6189ab73003b54b632aa39679016d Mon Sep 17 00:00:00 2001 From: kckennylau Date: Sat, 23 Feb 2019 01:10:18 +0000 Subject: [PATCH 13/22] use sqrt 2 instead of F4 --- src/ring_theory/free_comm_ring_support.lean | 65 ++------------------- 1 file changed, 6 insertions(+), 59 deletions(-) diff --git a/src/ring_theory/free_comm_ring_support.lean b/src/ring_theory/free_comm_ring_support.lean index 80762389ffc9c..19962da38dad6 100644 --- a/src/ring_theory/free_comm_ring_support.lean +++ b/src/ring_theory/free_comm_ring_support.lean @@ -1,56 +1,5 @@ import ring_theory.free_comm_ring ring_theory.subring data.set.finite data.real.irrational -@[derive decidable_eq] -inductive F4 : Type | z | o | a | b - -namespace F4 - -instance : fintype F4 := -{ elems := {z, o, a, b}, - complete := λ x, by cases x; simp } - -def add : F4 → F4 → F4 -| z x := x -| x z := x -| o o := z -| o a := b -| o b := a -| a o := b -| a a := z -| a b := o -| b o := a -| b a := o -| b b := z - -def mul : F4 → F4 → F4 -| z x := z -| x z := z -| o x := x -| x o := x -| a a := b -| a b := o -| b a := o -| b b := a - -def inv : F4 → F4 -| z := z -| o := o -| a := b -| b := a - -instance : discrete_field F4 := -by refine -{ add := add, - zero := z, - neg := id, - mul := mul, - one := o, - inv := inv, - has_decidable_eq := F4.decidable_eq, - .. }; tactic.exact_dec_trivial - -end F4 - universes u v variables {α : Type u} @@ -110,16 +59,14 @@ theorem is_supported_pure {p} {s : set α} : is_supported (pure p) s ↔ p ∈ s suffices is_supported (pure p) s → p ∈ s, from ⟨this, λ hps, ring.subset_closure ⟨p, hps, rfl⟩⟩, assume hps : is_supported (pure p) s, begin classical, - have : ∀ x, is_supported x s → lift (λ q, if q ∈ s then 0 else F4.a) x = 0 ∨ - lift (λ q, if q ∈ s then 0 else F4.a) x = 1, + have : ∀ x, is_supported x s → ∃ q : ℚ, lift (λ q, if q ∈ s then 0 else real.sqrt 2) x = ↑q, { intros x hx, refine ring.in_closure.rec_on hx _ _ _ _, - { right, rw lift_one }, - { right, rw [lift_neg, lift_one], refl }, - { rintros _ ⟨q, hqs, rfl⟩ _ _, left, rw [lift_mul, lift_of, if_pos hqs, zero_mul] }, - { rintros x y (ihx | ihx) (ihy | ihy); rw [lift_add, ihx, ihy], - exacts [or.inl rfl, or.inr rfl, or.inr rfl, or.inl rfl] } }, + { use 1, rw [lift_one, rat.cast_one] }, + { use -1, rw [lift_neg, lift_one, rat.cast_neg, rat.cast_one], }, + { rintros _ ⟨z, hzs, rfl⟩ _ _, use 0, rw [lift_mul, lift_of, if_pos hzs, zero_mul, rat.cast_zero] }, + { rintros x y ⟨q, hq⟩ ⟨r, hr⟩, use q+r, rw [lift_add, hq, hr, rat.cast_add] } }, specialize this (of p) hps, rw [lift_of] at this, split_ifs at this, { exact h }, - exfalso, cases this; cases this + exfalso, exact irr_sqrt_two this end theorem subtype_val_map_restriction {x} (s : set α) [decidable_pred s] (hxs : is_supported x s) : From cf8a619bde24d327387dd7321a8da55fd3bbd585 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Sat, 23 Feb 2019 10:30:28 +0000 Subject: [PATCH 14/22] direct limit of field --- src/algebra/direct_limit.lean | 157 +++++++++++++------- src/ring_theory/free_comm_ring.lean | 86 ++++++++++- src/ring_theory/free_comm_ring_support.lean | 93 ------------ 3 files changed, 186 insertions(+), 150 deletions(-) delete mode 100644 src/ring_theory/free_comm_ring_support.lean diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 7d2139afce0e1..fcfd7ab305947 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -1,6 +1,6 @@ import linear_algebra.direct_sum_module import linear_algebra.linear_combination -import ring_theory.free_comm_ring_support +import ring_theory.free_comm_ring import ring_theory.ideal_operations import ring_theory.to_field @@ -11,6 +11,35 @@ open lattice submodule class directed_order (α : Type u) extends partial_order α := (directed : ∀ i j : α, ∃ k, i ≤ k ∧ j ≤ k) +theorem finset.exists_le_of_nonempty {α : Type u} [directed_order α] (s : finset α) (hs : s ≠ ∅) : + ∃ M, ∀ i ∈ s, i ≤ M := +begin + revert hs, classical, refine finset.induction_on s (λ H, false.elim $ H rfl) _, + intros a s has ih _, by_cases hs : s = ∅, + { rw hs, exact ⟨a, λ i hi, finset.mem_singleton.1 hi ▸ le_refl i⟩ }, + rcases ih hs with ⟨j, hj⟩, rcases directed_order.directed a j with ⟨k, hak, hjk⟩, + exact ⟨k, λ i hi, or.cases_on (finset.mem_insert.1 hi) (λ hia, hia.symm ▸ hak) (λ his, le_trans (hj i his) hjk)⟩ +end + +theorem finset.exists_le {α : Type u} [nonempty α] [directed_order α] (s : finset α) : + ∃ M, ∀ i ∈ s, i ≤ M := +nonempty.elim (by apply_instance) $ assume ind : α, +classical.by_cases + (assume hs : s = ∅, ⟨ind, hs.symm ▸ λ i, false.elim⟩) + s.exists_le_of_nonempty + +theorem set.finite.exists_le_of_nonempty {α : Type u} [directed_order α] {s : set α} (hs : s ≠ ∅) (hfs : set.finite s) : + ∃ M, ∀ i ∈ s, i ≤ M := +let ⟨x, hxs⟩ := set.ne_empty_iff_exists_mem.1 hs in +have hfs.to_finset ≠ ∅, from finset.ne_empty_of_mem (set.finite.mem_to_finset.2 hxs), +let ⟨M, HM⟩ := hfs.to_finset.exists_le_of_nonempty this in +⟨M, λ i his, HM i $ set.finite.mem_to_finset.2 his⟩ + +theorem set.finite.exists_le {α : Type u} [nonempty α] [directed_order α] {s : set α} (hfs : set.finite s) : + ∃ M, ∀ i ∈ s, i ≤ M := +let ⟨M, HM⟩ := hfs.to_finset.exists_le in +⟨M, λ i his, HM i $ set.finite.mem_to_finset.2 his⟩ + variables {R : Type u} [ring R] variables {ι : Type v} [nonempty ι] variables [directed_order ι] [decidable_eq ι] @@ -295,21 +324,29 @@ instance of.is_ring_hom (i) : is_ring_hom (of G f i) := map_add := λ x y, ideal.quotient.eq.2 $ subset_span $ or.inl $ or.inr $ set.mem_Union.2 ⟨i, set.mem_Union.2 ⟨x, set.mem_range_self y⟩⟩ } -theorem of_f {i j} (hij : i ≤ j) (x : G i) : of G f j (f i j hij x) = of G f i x := +@[simp] lemma of_f {i j} (hij : i ≤ j) (x : G i) : of G f j (f i j hij x) = of G f i x := ideal.quotient.eq.2 $ subset_span $ or.inl $ or.inl $ or.inl $ set.mem_Union.2 ⟨i, set.mem_Union.2 ⟨j, set.mem_Union.2 ⟨hij, set.mem_range_self _⟩⟩⟩ +@[simp] lemma of_zero (i) : of G f i 0 = 0 := is_ring_hom.map_zero _ +@[simp] lemma of_one (i) : of G f i 1 = 1 := is_ring_hom.map_one _ +@[simp] lemma of_add (i x y) : of G f i (x + y) = of G f i x + of G f i y := is_ring_hom.map_add _ +@[simp] lemma of_neg (i x) : of G f i (-x) = -of G f i x := is_ring_hom.map_neg _ +@[simp] lemma of_sub (i x y) : of G f i (x - y) = of G f i x - of G f i y := is_ring_hom.map_sub _ +@[simp] lemma of_mul (i x y) : of G f i (x * y) = of G f i x * of G f i y := is_ring_hom.map_mul _ +@[simp] lemma of_pow (i x) (n : ℕ) : of G f i (x ^ n) = of G f i x ^ n := is_semiring_hom.map_pow _ _ _ + theorem exists_of (z : direct_limit G f) : ∃ i x, of G f i x = z := nonempty.elim (by apply_instance) $ assume ind : ι, quotient.induction_on' z $ λ x, free_abelian_group.induction_on x - ⟨ind, 0, is_ring_hom.map_zero (of G f ind)⟩ + ⟨ind, 0, of_zero G f ind⟩ (λ s, multiset.induction_on s - ⟨ind, 1, is_ring_hom.map_one (of G f ind)⟩ + ⟨ind, 1, of_one G f ind⟩ (λ a s ih, let ⟨i, x⟩ := a, ⟨j, y, hs⟩ := ih, ⟨k, hik, hjk⟩ := directed_order.directed i j in - ⟨k, f i k hik x * f j k hjk y, by rw [is_ring_hom.map_mul (of G f k), of_f, of_f, hs]; refl⟩)) - (λ s ⟨i, x, ih⟩, ⟨i, -x, by rw [is_ring_hom.map_neg (of G f i), ih]; refl⟩) + ⟨k, f i k hik x * f j k hjk y, by rw [of_mul, of_f, of_f, hs]; refl⟩)) + (λ s ⟨i, x, ih⟩, ⟨i, -x, by rw [of_neg, ih]; refl⟩) (λ p q ⟨i, x, ihx⟩ ⟨j, y, ihy⟩, let ⟨k, hik, hjk⟩ := directed_order.directed i j in - ⟨k, f i k hik x + f j k hjk y, by rw [is_ring_hom.map_add (of G f k), of_f, of_f, ihx, ihy]; refl⟩) + ⟨k, f i k hik x + f j k hjk y, by rw [of_add, of_f, of_f, ihx, ihy]; refl⟩) @[elab_as_eliminator] theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit G f) (ih : ∀ i x, C (of G f i x)) : C z := @@ -359,36 +396,27 @@ variables [add_comm_group.directed_system G f] open free_comm_ring -/-section totalize -local attribute [instance, priority 0] classical.dec - -noncomputable def totalize : Π i j, G i → G j := -λ i j, if h : i ≤ j then f i j h else 0 - -instance totalize.is_ring_hom (i j) : is_ring_hom (totalize G f i j) := -by unfold totalize; split_ifs; apply_instance - -lemma totalize_apply (i j x) : - totalize G f i j x = if h : i ≤ j then f i j h x else 0 := -if h : i ≤ j - then by dsimp only [totalize]; rw [dif_pos h, dif_pos h] - else by dsimp only [totalize]; rw [dif_neg h, dif_neg h, pi.zero_apply] -end totalize-/ - -/-lemma to_module_totalize_of_le {x : free_comm_ring Σ i, G i} {i j : ι} - (hij : i ≤ j) (hx : ∀ k ∈ x.support, k ≤ i) : - direct_sum.to_module R ι (G j) (λ k, totalize G f k j) x = - f i j hij (direct_sum.to_module R ι (G i) (λ k, totalize G f k i) x) := -begin - rw [← @dfinsupp.sum_single ι G _ _ _ x, dfinsupp.sum], - simp only [linear_map.map_sum], - refine finset.sum_congr rfl (λ k hk, _), - rw direct_sum.single_eq_lof R k (x k), - simp [totalize_apply, hx k hk, le_trans (hx k hk) hij, directed_system.Hcomp f] -end-/ - section of_zero_exact_aux attribute [instance, priority 0] classical.dec +lemma of.zero_exact_aux2 {x : free_comm_ring Σ i, G i} {s t} (hxs : is_supported x s) {j k} + (hj : ∀ z : Σ i, G i, z ∈ s → z.1 ≤ j) (hk : ∀ z : Σ i, G i, z ∈ t → z.1 ≤ k) + (hjk : j ≤ k) (hst : s ⊆ t) : + f j k hjk (lift (λ ix : s, f ix.1.1 j (hj ix ix.2) ix.1.2) (restriction s x)) = + lift (λ ix : t, f ix.1.1 k (hk ix ix.2) ix.1.2) (restriction t x) := +begin + refine ring.in_closure.rec_on hxs _ _ _ _, + { rw [restriction_one, lift_one, is_ring_hom.map_one (f j k hjk), restriction_one, lift_one] }, + { rw [restriction_neg, restriction_one, lift_neg, lift_one, + is_ring_hom.map_neg (f j k hjk), is_ring_hom.map_one (f j k hjk), + restriction_neg, restriction_one, lift_neg, lift_one] }, + { rintros _ ⟨p, hps, rfl⟩ n ih, + rw [restriction_mul, lift_mul, is_ring_hom.map_mul (f j k hjk), ih, restriction_mul, lift_mul, + restriction_of, dif_pos hps, lift_pure, restriction_of, dif_pos (hst hps), lift_pure], + dsimp only, rw add_comm_group.directed_system.Hcomp f, refl }, + { rintros x y ihx ihy, + rw [restriction_add, lift_add, is_ring_hom.map_add (f j k hjk), ihx, ihy, restriction_add, lift_add] } +end + lemma of.zero_exact_aux {x : free_comm_ring Σ i, G i} (H : x ∈ ideal.span (thing1 G f ∪ thing2 G ∪ thing3 G ∪ thing4 G)) : ∃ j s, ∃ H : (∀ k : Σ i, G i, k ∈ s → k.1 ≤ j), is_supported x s ∧ lift (λ ix : s, f ix.1.1 j (H ix ix.2) ix.1.2) (restriction s x) = (0 : G j) := @@ -414,8 +442,8 @@ begin { rintros k (rfl | ⟨rfl | ⟨rfl | hk⟩⟩), refl, refl, refl, cases hk }, { rw [restriction_sub, restriction_add, restriction_of, restriction_of, restriction_of, dif_pos, dif_pos, dif_pos, lift_sub, lift_add, lift_pure, lift_pure, lift_pure], - dsimp only, simp only [add_comm_group.directed_system.Hid f], rw sub_self, - { left, refl }, { right, left, refl }, { right, right, left, refl } } }, + dsimp only, rw is_ring_hom.map_add (f i i _), exact sub_self _, + { right, right, left, refl }, { apply_instance }, { left, refl }, { right, left, refl } } }, { simp only [thing4, set.mem_Union, set.mem_range] at hx, rcases hx with ⟨i, x, y, rfl⟩, refine ⟨i, {⟨i, x*y⟩, ⟨i, x⟩, ⟨i, y⟩}, _, is_supported_sub (is_supported_pure.2 $ or.inr $ or.inr $ or.inl rfl) @@ -423,18 +451,32 @@ begin { rintros k (rfl | ⟨rfl | ⟨rfl | hk⟩⟩), refl, refl, refl, cases hk }, { rw [restriction_sub, restriction_mul, restriction_of, restriction_of, restriction_of, dif_pos, dif_pos, dif_pos, lift_sub, lift_mul, lift_pure, lift_pure, lift_pure], - dsimp only, simp only [add_comm_group.directed_system.Hid f], rw sub_self, - { left, refl }, { right, left, refl }, { right, right, left, refl } } } }, + dsimp only, rw is_ring_hom.map_mul (f i i _), exact sub_self _, + { right, right, left, refl }, { apply_instance }, { left, refl }, { right, left, refl } } } }, { refine nonempty.elim (by apply_instance) (assume ind : ι, _), refine ⟨ind, ∅, λ _, false.elim, is_supported_zero, _⟩, rw [restriction_zero, lift_zero] }, - { rintros x y ⟨i, s, hi, hxs, ihs⟩ ⟨j, t, hj, hxt, iht⟩, + { rintros x y ⟨i, s, hi, hxs, ihs⟩ ⟨j, t, hj, hyt, iht⟩, + rcases directed_order.directed i j with ⟨k, hik, hjk⟩, + have : ∀ z : Σ i, G i, z ∈ s ∪ t → z.1 ≤ k, + { rintros z (hz | hz), exact le_trans (hi z hz) hik, exact le_trans (hj z hz) hjk }, + refine ⟨k, s ∪ t, this, is_supported_add (is_supported_upwards hxs $ set.subset_union_left s t) + (is_supported_upwards hyt $ set.subset_union_right s t), _⟩, + { rw [restriction_add, lift_add, + ← of.zero_exact_aux2 G f hxs hi this hik (set.subset_union_left s t), + ← of.zero_exact_aux2 G f hyt hj this hjk (set.subset_union_right s t), + ihs, is_ring_hom.map_zero (f i k hik), iht, is_ring_hom.map_zero (f j k hjk), zero_add] } }, + { rintros x y ⟨j, t, hj, hyt, iht⟩, rw smul_eq_mul, + rcases exists_finset x with ⟨s, hfs, hxs⟩, + rcases (set.finite_image sigma.fst hfs).exists_le with ⟨i, hi⟩, rcases directed_order.directed i j with ⟨k, hik, hjk⟩, - refine ⟨k, s ∪ t, _, is_supported_add (is_supported_upwards hxs $ set.subset_union_left s t) - (is_supported_upwards hxt $ set.subset_union_right s t), _⟩, - { rintros t (ht | ht), exact le_trans (hi _ ht) hik, exact le_trans (hj _ ht) hjk }, - { rw [restriction_add, lift_add], sorry } }, - sorry + have : ∀ z : Σ i, G i, z ∈ s ∪ t → z.1 ≤ k, + { rintros z (hz | hz), exact le_trans (hi z.1 ⟨z, hz, rfl⟩) hik, exact le_trans (hj z hz) hjk }, + refine ⟨k, s ∪ t, this, is_supported_mul (is_supported_upwards hxs $ set.subset_union_left s t) + (is_supported_upwards hyt $ set.subset_union_right s t), _⟩, + rw [restriction_mul, lift_mul, + ← of.zero_exact_aux2 G f hyt hj this hjk (set.subset_union_right s t), + iht, is_ring_hom.map_zero (f j k hjk), mul_zero] } end end of_zero_exact_aux @@ -442,19 +484,16 @@ lemma of.zero_exact {i x} (hix : of G f i x = 0) : ∃ j, ∃ hij : i ≤ j, f i let ⟨j, s, H, hxs, hx⟩ := of.zero_exact_aux G f (ideal.quotient.eq_zero_iff_mem.1 hix) in have hixs : (⟨i, x⟩ : Σ i, G i) ∈ s, from is_supported_pure.1 hxs, ⟨j, H ⟨i, x⟩ hixs, by rw [restriction_of, dif_pos hixs, lift_pure] at hx; exact hx⟩ -/- + theorem of_inj (hf : ∀ i j hij, function.injective (f i j hij)) (i) : function.injective (of G f i) := begin suffices : ∀ x, of G f i x = 0 → x = 0, { intros x y hxy, rw ← sub_eq_zero_iff_eq, apply this, rw [is_ring_hom.map_sub (of G f i), hxy, sub_self] }, - intros x hx, replace hx := ideal.quotient.eq_zero_iff_mem.1 hx, - generalize_hyp hy : free_comm_ring.of _ = y at hx, revert hy, - refine span_induction hx _ _ _ _, - { rintros z ⟨⟨hz | hz⟩ | hz⟩, - { simp only [thing1, set.mem_Union, set.mem_range] at hz } } -end-/ + intros x hx, rcases of.zero_exact G f hx with ⟨j, hij, hfx⟩, + apply hf i j hij, rw [hfx, is_ring_hom.map_zero (f i j hij)] +end end direct_limit @@ -468,14 +507,20 @@ variables (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_field_hom (f i j variables [add_comm_group.directed_system G f] instance direct_limit.nonzero_comm_ring : nonzero_comm_ring (ring.direct_limit G f) := -{ zero_ne_one := sorry, +{ zero_ne_one := nonempty.elim (by apply_instance) $ assume i : ι, begin + change (0 : ring.direct_limit G f) ≠ 1, + rw ← ring.direct_limit.of_one, + intros H, rcases ring.direct_limit.of.zero_exact G f H.symm with ⟨j, hij, hf⟩, + rw is_ring_hom.map_one (f i j hij) at hf, + exact one_ne_zero hf + end, .. ring.direct_limit.comm_ring G f } instance direct_limit.is_division_ring : is_division_ring (ring.direct_limit G f) := { exists_inv := λ p, ring.direct_limit.induction_on G f p $ λ i x H, - ⟨ring.direct_limit.of G f i (x⁻¹), by erw [← is_ring_hom.map_mul (ring.direct_limit.of G f i), - mul_inv_cancel (assume h : x = 0, H $ by rw [h, is_ring_hom.map_zero (ring.direct_limit.of G f i)]), - is_ring_hom.map_one (ring.direct_limit.of G f i)]⟩ } + ⟨ring.direct_limit.of G f i (x⁻¹), by erw [← ring.direct_limit.of_mul, + mul_inv_cancel (assume h : x = 0, H $ by rw [h, ring.direct_limit.of_zero]), + ring.direct_limit.of_one]⟩ } def direct_limit : Type (max v w) := to_field (ring.direct_limit G f) diff --git a/src/ring_theory/free_comm_ring.lean b/src/ring_theory/free_comm_ring.lean index 12775debff2ae..6a418e15cea5a 100644 --- a/src/ring_theory/free_comm_ring.lean +++ b/src/ring_theory/free_comm_ring.lean @@ -1,4 +1,4 @@ -import group_theory.free_abelian_group +import group_theory.free_abelian_group ring_theory.subring data.set.finite data.real.irrational universes u v @@ -174,4 +174,88 @@ variables {β : Type u} (f : α → β) (x y : free_comm_ring α) @[simp] lemma map_mul : f <$> (x * y) = f <$> x * f <$> y := lift_mul (of ∘ f) x y end map +def is_supported (x : free_comm_ring α) (s : set α) : Prop := +x ∈ ring.closure (of '' s) + +theorem is_supported_upwards {x : free_comm_ring α} {s t} (hs : is_supported x s) (hst : s ⊆ t) : + is_supported x t := +ring.closure_mono (set.mono_image hst) hs + +theorem is_supported_add {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : + is_supported (x + y) s := +is_add_submonoid.add_mem hxs hys + +theorem is_supported_neg {x : free_comm_ring α} {s} (hxs : is_supported x s) : + is_supported (-x) s := +is_add_subgroup.neg_mem hxs + +theorem is_supported_sub {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : + is_supported (x - y) s := +is_add_subgroup.sub_mem _ _ _ hxs hys + +theorem is_supported_mul {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : + is_supported (x * y) s := +is_submonoid.mul_mem hxs hys + +theorem is_supported_zero {s : set α} : is_supported 0 s := +is_add_submonoid.zero_mem _ + +theorem is_supported_one {s : set α} : is_supported 1 s := +is_submonoid.one_mem _ + +theorem is_supported_int {i : ℤ} {s : set α} : is_supported ↑i s := +int.induction_on i is_supported_zero + (λ i hi, by rw [int.cast_add, int.cast_one]; exact is_supported_add hi is_supported_one) + (λ i hi, by rw [int.cast_sub, int.cast_one]; exact is_supported_sub hi is_supported_one) + +def restriction (s : set α) [decidable_pred s] (x : free_comm_ring α) : free_comm_ring s := +lift (λ p, if H : p ∈ s then pure ⟨p, H⟩ else 0) x + +section restriction +variables (s : set α) [decidable_pred s] (x y : free_comm_ring α) +@[simp] lemma restriction_of (p) : restriction s (of p) = if H : p ∈ s then pure ⟨p, H⟩ else 0 := lift_of _ _ +@[simp] lemma restriction_pure (p) : restriction s (pure p) = if H : p ∈ s then pure ⟨p, H⟩ else 0 := lift_of _ _ +@[simp] lemma restriction_zero : restriction s 0 = 0 := lift_zero _ +@[simp] lemma restriction_one : restriction s 1 = 1 := lift_one _ +@[simp] lemma restriction_add : restriction s (x + y) = restriction s x + restriction s y := lift_add _ _ _ +@[simp] lemma restriction_neg : restriction s (-x) = -restriction s x := lift_neg _ _ +@[simp] lemma restriction_sub : restriction s (x - y) = restriction s x - restriction s y := lift_sub _ _ _ +@[simp] lemma restriction_mul : restriction s (x * y) = restriction s x * restriction s y := lift_mul _ _ _ +end restriction + +theorem is_supported_pure {p} {s : set α} : is_supported (pure p) s ↔ p ∈ s := +suffices is_supported (pure p) s → p ∈ s, from ⟨this, λ hps, ring.subset_closure ⟨p, hps, rfl⟩⟩, +assume hps : is_supported (pure p) s, begin + classical, + have : ∀ x, is_supported x s → ∃ q : ℚ, lift (λ q, if q ∈ s then 0 else real.sqrt 2) x = ↑q, + { intros x hx, refine ring.in_closure.rec_on hx _ _ _ _, + { use 1, rw [lift_one, rat.cast_one] }, + { use -1, rw [lift_neg, lift_one, rat.cast_neg, rat.cast_one], }, + { rintros _ ⟨z, hzs, rfl⟩ _ _, use 0, rw [lift_mul, lift_of, if_pos hzs, zero_mul, rat.cast_zero] }, + { rintros x y ⟨q, hq⟩ ⟨r, hr⟩, use q+r, rw [lift_add, hq, hr, rat.cast_add] } }, + specialize this (of p) hps, rw [lift_of] at this, split_ifs at this, { exact h }, + exfalso, exact irr_sqrt_two this +end + +theorem subtype_val_map_restriction {x} (s : set α) [decidable_pred s] (hxs : is_supported x s) : + subtype.val <$> restriction s x = x := +begin + refine ring.in_closure.rec_on hxs _ _ _ _, + { rw restriction_one, refl }, + { rw [restriction_neg, map_neg, restriction_one], refl }, + { rintros _ ⟨p, hps, rfl⟩ n ih, rw [restriction_mul, restriction_of, dif_pos hps, map_mul, map_pure, ih], refl }, + { intros x y ihx ihy, rw [restriction_add, map_add, ihx, ihy] } +end + +theorem exists_finset (x : free_comm_ring α) : ∃ s : set α, set.finite s ∧ is_supported x s := +free_comm_ring.induction_on x + ⟨∅, set.finite_empty, is_supported_neg is_supported_one⟩ + (λ p, ⟨{p}, set.finite_singleton p, is_supported_pure.2 $ finset.mem_singleton_self _⟩) + (λ x y ⟨s, hfs, hxs⟩ ⟨t, hft, hxt⟩, ⟨s ∪ t, set.finite_union hfs hft, is_supported_add + (is_supported_upwards hxs $ set.subset_union_left s t) + (is_supported_upwards hxt $ set.subset_union_right s t)⟩) + (λ x y ⟨s, hfs, hxs⟩ ⟨t, hft, hxt⟩, ⟨s ∪ t, set.finite_union hfs hft, is_supported_mul + (is_supported_upwards hxs $ set.subset_union_left s t) + (is_supported_upwards hxt $ set.subset_union_right s t)⟩) + end free_comm_ring diff --git a/src/ring_theory/free_comm_ring_support.lean b/src/ring_theory/free_comm_ring_support.lean deleted file mode 100644 index 19962da38dad6..0000000000000 --- a/src/ring_theory/free_comm_ring_support.lean +++ /dev/null @@ -1,93 +0,0 @@ -import ring_theory.free_comm_ring ring_theory.subring data.set.finite data.real.irrational - -universes u v - -variables {α : Type u} - -namespace free_comm_ring - -def is_supported (x : free_comm_ring α) (s : set α) : Prop := -x ∈ ring.closure (of '' s) - -theorem is_supported_upwards {x : free_comm_ring α} {s t} (hs : is_supported x s) (hst : s ⊆ t) : - is_supported x t := -ring.closure_mono (set.mono_image hst) hs - -theorem is_supported_add {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : - is_supported (x + y) s := -is_add_submonoid.add_mem hxs hys - -theorem is_supported_neg {x : free_comm_ring α} {s} (hxs : is_supported x s) : - is_supported (-x) s := -is_add_subgroup.neg_mem hxs - -theorem is_supported_sub {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : - is_supported (x - y) s := -is_add_subgroup.sub_mem _ _ _ hxs hys - -theorem is_supported_mul {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : - is_supported (x * y) s := -is_submonoid.mul_mem hxs hys - -theorem is_supported_zero {s : set α} : is_supported 0 s := -is_add_submonoid.zero_mem _ - -theorem is_supported_one {s : set α} : is_supported 1 s := -is_submonoid.one_mem _ - -theorem is_supported_int {i : ℤ} {s : set α} : is_supported ↑i s := -int.induction_on i is_supported_zero - (λ i hi, by rw [int.cast_add, int.cast_one]; exact is_supported_add hi is_supported_one) - (λ i hi, by rw [int.cast_sub, int.cast_one]; exact is_supported_sub hi is_supported_one) - -def restriction (s : set α) [decidable_pred s] (x : free_comm_ring α) : free_comm_ring s := -lift (λ p, if H : p ∈ s then pure ⟨p, H⟩ else 0) x - -section restriction -variables (s : set α) [decidable_pred s] (x y : free_comm_ring α) -@[simp] lemma restriction_of (p) : restriction s (of p) = if H : p ∈ s then pure ⟨p, H⟩ else 0 := lift_of _ _ -@[simp] lemma restriction_pure (p) : restriction s (pure p) = if H : p ∈ s then pure ⟨p, H⟩ else 0 := lift_of _ _ -@[simp] lemma restriction_zero : restriction s 0 = 0 := lift_zero _ -@[simp] lemma restriction_one : restriction s 1 = 1 := lift_one _ -@[simp] lemma restriction_add : restriction s (x + y) = restriction s x + restriction s y := lift_add _ _ _ -@[simp] lemma restriction_neg : restriction s (-x) = -restriction s x := lift_neg _ _ -@[simp] lemma restriction_sub : restriction s (x - y) = restriction s x - restriction s y := lift_sub _ _ _ -@[simp] lemma restriction_mul : restriction s (x * y) = restriction s x * restriction s y := lift_mul _ _ _ -end restriction - -theorem is_supported_pure {p} {s : set α} : is_supported (pure p) s ↔ p ∈ s := -suffices is_supported (pure p) s → p ∈ s, from ⟨this, λ hps, ring.subset_closure ⟨p, hps, rfl⟩⟩, -assume hps : is_supported (pure p) s, begin - classical, - have : ∀ x, is_supported x s → ∃ q : ℚ, lift (λ q, if q ∈ s then 0 else real.sqrt 2) x = ↑q, - { intros x hx, refine ring.in_closure.rec_on hx _ _ _ _, - { use 1, rw [lift_one, rat.cast_one] }, - { use -1, rw [lift_neg, lift_one, rat.cast_neg, rat.cast_one], }, - { rintros _ ⟨z, hzs, rfl⟩ _ _, use 0, rw [lift_mul, lift_of, if_pos hzs, zero_mul, rat.cast_zero] }, - { rintros x y ⟨q, hq⟩ ⟨r, hr⟩, use q+r, rw [lift_add, hq, hr, rat.cast_add] } }, - specialize this (of p) hps, rw [lift_of] at this, split_ifs at this, { exact h }, - exfalso, exact irr_sqrt_two this -end - -theorem subtype_val_map_restriction {x} (s : set α) [decidable_pred s] (hxs : is_supported x s) : - subtype.val <$> restriction s x = x := -begin - refine ring.in_closure.rec_on hxs _ _ _ _, - { rw restriction_one, refl }, - { rw [restriction_neg, map_neg, restriction_one], refl }, - { rintros _ ⟨p, hps, rfl⟩ n ih, rw [restriction_mul, restriction_of, dif_pos hps, map_mul, map_pure, ih], refl }, - { intros x y ihx ihy, rw [restriction_add, map_add, ihx, ihy] } -end - -theorem exists_finset (x : free_comm_ring α) : ∃ s : set α, set.finite s ∧ is_supported x s := -free_comm_ring.induction_on x - ⟨∅, set.finite_empty, is_supported_neg is_supported_one⟩ - (λ p, ⟨{p}, set.finite_singleton p, is_supported_pure.2 $ finset.mem_singleton_self _⟩) - (λ x y ⟨s, hfs, hxs⟩ ⟨t, hft, hxt⟩, ⟨s ∪ t, set.finite_union hfs hft, is_supported_add - (is_supported_upwards hxs $ set.subset_union_left s t) - (is_supported_upwards hxt $ set.subset_union_right s t)⟩) - (λ x y ⟨s, hfs, hxs⟩ ⟨t, hft, hxt⟩, ⟨s ∪ t, set.finite_union hfs hft, is_supported_mul - (is_supported_upwards hxs $ set.subset_union_left s t) - (is_supported_upwards hxt $ set.subset_union_right s t)⟩) - -end free_comm_ring From 5c0258ae27f5462a9e1ed0228c47ea567f709f13 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Sat, 23 Feb 2019 17:10:12 +0000 Subject: [PATCH 15/22] cleanup for mathlib --- src/algebra/big_operators.lean | 4 + src/algebra/direct_limit.lean | 450 ++++++++++++---------------- src/algebra/module.lean | 15 + src/linear_algebra/basic.lean | 7 + src/order/basic.lean | 3 + src/ring_theory/free_comm_ring.lean | 11 +- src/ring_theory/to_field.lean | 231 -------------- 7 files changed, 228 insertions(+), 493 deletions(-) delete mode 100644 src/ring_theory/to_field.lean diff --git a/src/algebra/big_operators.lean b/src/algebra/big_operators.lean index 0afc0a8166af1..6dfde585e9f17 100644 --- a/src/algebra/big_operators.lean +++ b/src/algebra/big_operators.lean @@ -21,6 +21,10 @@ multiset.induction_on s.1 (let ⟨z⟩ := hι in ⟨z, λ _, false.elim⟩) $ (λ h, h.symm ▸ h₁) (λ h, trans (H _ h) h₂)⟩ +theorem finset.exists_le {α : Type u} [nonempty α] [directed_order α] (s : finset α) : + ∃ M, ∀ i ∈ s, i ≤ M := +directed.finset_le (by apply_instance) directed_order.directed s + namespace finset variables {s s₁ s₂ : finset α} {a : α} {f g : α → β} diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index fcfd7ab305947..606aab7b1ee02 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -1,50 +1,29 @@ +/- +Copyright (c) 2019 Kenny Lau, Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau, Chris Hughes + +Direct limit of modules, abelian groups, rings, and fields. +-/ + import linear_algebra.direct_sum_module import linear_algebra.linear_combination import ring_theory.free_comm_ring import ring_theory.ideal_operations -import ring_theory.to_field universes u v w u₁ open lattice submodule -class directed_order (α : Type u) extends partial_order α := -(directed : ∀ i j : α, ∃ k, i ≤ k ∧ j ≤ k) - -theorem finset.exists_le_of_nonempty {α : Type u} [directed_order α] (s : finset α) (hs : s ≠ ∅) : - ∃ M, ∀ i ∈ s, i ≤ M := -begin - revert hs, classical, refine finset.induction_on s (λ H, false.elim $ H rfl) _, - intros a s has ih _, by_cases hs : s = ∅, - { rw hs, exact ⟨a, λ i hi, finset.mem_singleton.1 hi ▸ le_refl i⟩ }, - rcases ih hs with ⟨j, hj⟩, rcases directed_order.directed a j with ⟨k, hak, hjk⟩, - exact ⟨k, λ i hi, or.cases_on (finset.mem_insert.1 hi) (λ hia, hia.symm ▸ hak) (λ his, le_trans (hj i his) hjk)⟩ -end - -theorem finset.exists_le {α : Type u} [nonempty α] [directed_order α] (s : finset α) : - ∃ M, ∀ i ∈ s, i ≤ M := -nonempty.elim (by apply_instance) $ assume ind : α, -classical.by_cases - (assume hs : s = ∅, ⟨ind, hs.symm ▸ λ i, false.elim⟩) - s.exists_le_of_nonempty - -theorem set.finite.exists_le_of_nonempty {α : Type u} [directed_order α] {s : set α} (hs : s ≠ ∅) (hfs : set.finite s) : - ∃ M, ∀ i ∈ s, i ≤ M := -let ⟨x, hxs⟩ := set.ne_empty_iff_exists_mem.1 hs in -have hfs.to_finset ≠ ∅, from finset.ne_empty_of_mem (set.finite.mem_to_finset.2 hxs), -let ⟨M, HM⟩ := hfs.to_finset.exists_le_of_nonempty this in -⟨M, λ i his, HM i $ set.finite.mem_to_finset.2 his⟩ - -theorem set.finite.exists_le {α : Type u} [nonempty α] [directed_order α] {s : set α} (hfs : set.finite s) : - ∃ M, ∀ i ∈ s, i ≤ M := -let ⟨M, HM⟩ := hfs.to_finset.exists_le in -⟨M, λ i his, HM i $ set.finite.mem_to_finset.2 his⟩ - variables {R : Type u} [ring R] variables {ι : Type v} [nonempty ι] variables [directed_order ι] [decidable_eq ι] variables (G : ι → Type w) [Π i, decidable_eq (G i)] +class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := +(Hid : ∀ i x h, f i i h x = x) +(Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) + namespace module variables [Π i, add_comm_group (G i)] [Π i, module R (G i)] @@ -55,17 +34,9 @@ class directed_system (f : Π i j, i ≤ j → G i →ₗ[R] G j) : Prop := variables (f : Π i j, i ≤ j → G i →ₗ[R] G j) [directed_system G f] -def thing : set (direct_sum ι G) := -⋃ i j H, set.range $ λ x, direct_sum.lof R ι G i x - direct_sum.lof R ι G j (f i j H x) - def direct_limit : Type (max v w) := -(span R $ thing G f).quotient - -variables {G f} -lemma mem_thing {a : direct_sum ι G} : a ∈ thing G f ↔ ∃ (i j) (H : i ≤ j) x, - direct_sum.lof R ι G i x - direct_sum.lof R ι G j (f i j H x) = a := -by simp only [thing, set.mem_Union, set.mem_range] -variables (G f) +(span R $ { a | ∃ (i j) (H : i ≤ j) x, + direct_sum.lof R ι G i x - direct_sum.lof R ι G j (f i j H x) = a }).quotient namespace direct_limit @@ -75,25 +46,34 @@ instance : module R (direct_limit G f) := quotient.module _ variables (R ι) def of (i) : G i →ₗ[R] direct_limit G f := (mkq _).comp $ direct_sum.lof R ι G i -variables {R ι} +variables {R ι G f} theorem of_f {i j hij x} : (of R ι G f j (f i j hij x)) = of R ι G f i x := -eq.symm $ (submodule.quotient.eq _).2 $ subset_span $ -set.mem_Union.2 ⟨i, set.mem_Union.2 ⟨j, set.mem_Union.2 ⟨hij, set.mem_range_self x⟩⟩⟩ +eq.symm $ (submodule.quotient.eq _).2 $ subset_span ⟨i, j, hij, x, rfl⟩ + +theorem exists_of (z : direct_limit G f) : ∃ i x, z = of R ι G f i x := +nonempty.elim (by apply_instance) $ assume ind : ι, +quotient.induction_on' z $ λ z, direct_sum.induction_on z + ⟨ind, 0, (linear_map.map_zero _).symm⟩ + (λ i x, ⟨i, x, rfl⟩) + (λ p q ⟨i, x, ihx⟩ ⟨j, y, ihy⟩, let ⟨k, hik, hjk⟩ := directed_order.directed i j in + ⟨k, f i k hik x + f j k hjk y, by rw [linear_map.map_add, of_f, of_f, ← ihx, ← ihy]; refl⟩) + +@[elab_as_eliminator] +protected theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit G f) + (ih : ∀ i x, C (of R ι G f i x)) : C z := +let ⟨i, x, h⟩ := exists_of z in h.symm ▸ ih i x variables {P : Type u₁} [add_comm_group P] [module R P] (g : Π i, G i →ₗ[R] P) variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) include Hg -variables (R ι) +variables (R ι G f) def rec : direct_limit G f →ₗ[R] P := liftq _ (direct_sum.to_module R ι P g) - (span_le.2 $ set.Union_subset $ λ i, - set.Union_subset $ λ j, set.Union_subset $ λ hij, - set.range_subset_iff.2 $ λ x, linear_map.sub_mem_ker_iff.2 $ - by rw [direct_sum.to_module_lof, direct_sum.to_module_lof, Hg]) - -variables {R ι} + (span_le.2 $ λ a ⟨i, j, hij, x, hx⟩, by rw [← hx, mem_coe, linear_map.sub_mem_ker_iff, + direct_sum.to_module_lof, direct_sum.to_module_lof, Hg]) +variables {R ι G f} omit Hg lemma rec_of {i} (x) : rec R ι G f g Hg (of R ι G f i x) = g i x := @@ -102,37 +82,14 @@ direct_sum.to_module_lof R _ _ theorem rec_unique (F : direct_limit G f →ₗ[R] P) (x) : F x = rec R ι G f (λ i, F.comp $ of R ι G f i) (λ i j hij x, by rw [linear_map.comp_apply, of_f]; refl) x := -quotient.induction_on' x $ λ y, direct_sum.induction_on y - ((linear_map.map_zero _).trans (linear_map.map_zero _).symm) - (λ i x, eq.symm $ rec_of _ _ _ _ _) - (λ x y ihx ihy, (linear_map.map_add F (quotient.mk' x) (quotient.mk' y)).trans $ - ihx.symm ▸ ihy.symm ▸ (linear_map.map_add _ _ _).symm) - -theorem exists_of (z : direct_limit G f) : ∃ i x, z = of R ι G f i x := -nonempty.elim (by apply_instance) $ assume ind : ι, -quotient.induction_on' z $ λ z, direct_sum.induction_on z - ⟨ind, 0, (linear_map.map_zero _).symm⟩ - (λ i x, ⟨i, x, rfl⟩) - (λ p q ⟨i, x, ihx⟩ ⟨j, y, ihy⟩, let ⟨k, hik, hjk⟩ := directed_order.directed i j in - ⟨k, f i k hik x + f j k hjk y, by rw [linear_map.map_add, of_f, of_f, ← ihx, ← ihy]; refl⟩) - -@[elab_as_eliminator] theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit G f) - (ih : ∀ i x, C (of R ι G f i x)) : C z := -let ⟨i, x, h⟩ := exists_of G f z in h.symm ▸ ih i x - -lemma span_preimage_le_comap_span {R M N: Type*} [ring R] [add_comm_group M] [module R M] - [add_comm_group N] [module R N] (f : N →ₗ[R] M) (s : set M) : span R (f ⁻¹' s) ≤ (span R s).comap f := -λ x h, span_induction h - (by simp only [set.preimage, set.mem_set_of_eq, mem_comap]; exact λ x h, subset_span h) - (zero_mem ((span R s).comap f)) - (λ _ _ hx hy, add_mem ((span R s).comap f) hx hy) - (λ _ _ h, smul_mem ((span R s).comap f) _ h) +direct_limit.induction_on x $ λ i x, by rw rec_of; refl section totalize local attribute [instance, priority 0] classical.dec - +variables (G f) noncomputable def totalize : Π i j, G i →ₗ[R] G j := λ i j, if h : i ≤ j then f i j h else 0 +variables {G f} lemma totalize_apply (i j x) : totalize G f i j x = if h : i ≤ j then f i j h x else 0 := @@ -153,12 +110,11 @@ begin simp [totalize_apply, hx k hk, le_trans (hx k hk) hij, directed_system.Hcomp f] end -lemma of.zero_exact_aux {x : direct_sum ι G} (H : x ∈ span R (thing G f)) : +lemma of.zero_exact_aux {x : direct_sum ι G} (H : submodule.quotient.mk x = (0 : direct_limit G f)) : ∃ j, (∀ k ∈ x.support, k ≤ j) ∧ direct_sum.to_module R ι (G j) (λ i, totalize G f i j) x = (0 : G j) := nonempty.elim (by apply_instance) $ assume ind : ι, -span_induction H - (λ x hx, let ⟨i, j, hij, y, hxy⟩ := mem_thing.1 hx in - let ⟨k, hik, hjk⟩ := directed_order.directed i j in +span_induction ((quotient.mk_eq_zero _).1 H) + (λ x ⟨i, j, hij, y, hxy⟩, let ⟨k, hik, hjk⟩ := directed_order.directed i j in ⟨k, begin clear_, subst hxy, @@ -180,78 +136,32 @@ span_induction H (λ hl, le_trans (hi _ hl) hik) (λ hl, le_trans (hj _ hl) hjk), by simp [linear_map.map_add, hxi, hyj, - to_module_totalize_of_le G f hik hi, - to_module_totalize_of_le G f hjk hj]⟩) + to_module_totalize_of_le hik hi, + to_module_totalize_of_le hjk hj]⟩) (λ a x ⟨i, hi, hxi⟩, ⟨i, λ k hk, hi k (dfinsupp.support_smul hk), by simp [linear_map.map_smul, hxi]⟩) theorem of.zero_exact {i x} (H : of R ι G f i x = 0) : ∃ j hij, f i j hij x = (0 : G j) := -let ⟨j, hj, hxj⟩ := of.zero_exact_aux G f - ((submodule.quotient.eq _).1 H) in +let ⟨j, hj, hxj⟩ := of.zero_exact_aux H in if hx0 : x = 0 then ⟨i, le_refl _, by simp [hx0]⟩ else have hij : i ≤ j, from hj _ $ by simp [direct_sum.apply_eq_component, hx0], ⟨j, hij, by simpa [totalize_apply, hij] using hxj⟩ -end direct_limit -@[extensionality] theorem ext {α : Type u} [ring α] {β : Type v} [add_comm_group β] : - Π (M₁ M₂ : module α β) - (H : ∀ r x, @has_scalar.smul α β M₁.to_has_scalar r x = @has_scalar.smul α β M₂.to_has_scalar r x), - M₁ = M₂ := -by rintro ⟨⟨⟨f⟩, _, _, _, _, _, _⟩⟩ ⟨⟨⟨g⟩, _, _, _, _, _, _⟩⟩ H; congr' 3; ext; apply H +end direct_limit end module -namespace add_comm_group - -variables {α : Type u} [add_comm_group α] - -protected def module : module ℤ α := module.of_core -{ smul := gsmul, - mul_smul := λ r s x, gsmul_mul x r s, - smul_add := λ r x y, gsmul_add x y r, - add_smul := λ r s x, add_gsmul x r s, - one_smul := one_gsmul } - -instance module.subsingleton : subsingleton (module ℤ α) := -begin - constructor, intros M₁ M₂, ext i x, refine int.induction_on i _ _ _, - { rw [zero_smul, zero_smul] }, - { intros i ih, rw [add_smul, add_smul, ih, one_smul, one_smul] }, - { intros i ih, rw [sub_smul, sub_smul, ih, one_smul, one_smul] } -end - -end add_comm_group - -local attribute [instance] add_comm_group.module - -namespace is_add_group_hom - -variables {α : Type u} {β : Type v} [add_comm_group α] [add_comm_group β] - -def to_linear_map (f : α → β) [is_add_group_hom f] : α →ₗ[ℤ] β := -{ to_fun := f, - add := is_add_group_hom.add f, - smul := λ i x, int.induction_on i (by rw [zero_smul, zero_smul, is_add_group_hom.zero f]) - (λ i ih, by rw [add_smul, add_smul, is_add_group_hom.add f, ih, one_smul, one_smul]) - (λ i ih, by rw [sub_smul, sub_smul, is_add_group_hom.sub f, ih, one_smul, one_smul]) } - -end is_add_group_hom - namespace add_comm_group variables [Π i, add_comm_group (G i)] -class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := -(Hid : ∀ i x h, f i i h x = x) -(Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) - -def direct_limit - (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] : Type* := +def direct_limit (f : Π i j, i ≤ j → G i → G j) + [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] : Type* := @module.direct_limit ℤ _ ι _ _ _ G _ _ _ (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) ⟨directed_system.Hid f, directed_system.Hcomp f⟩ @@ -261,7 +171,7 @@ namespace direct_limit variables (f : Π i j, i ≤ j → G i → G j) variables [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] -def directed_system : module.directed_system G (λ (i j : ι) (hij : i ≤ j), is_add_group_hom.to_linear_map (f i j hij)) := +def directed_system : module.directed_system G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) := ⟨directed_system.Hid f, directed_system.Hcomp f⟩ local attribute [instance] directed_system @@ -269,42 +179,74 @@ local attribute [instance] directed_system instance : add_comm_group (direct_limit G f) := module.direct_limit.add_comm_group G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) +set_option class.instance_max_depth 50 +def of (i) : G i → direct_limit G f := +module.direct_limit.of ℤ ι G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) i +variables {G f} + +instance of.is_add_group_hom (i) : is_add_group_hom (of G f i) := +linear_map.is_add_group_hom _ + +@[simp] lemma of_f {i j} (hij) (x) : of G f j (f i j hij x) = of G f i x := +module.direct_limit.of_f + +@[simp] lemma of_zero (i) : of G f i 0 = 0 := is_add_group_hom.zero _ +@[simp] lemma of_add (i x y) : of G f i (x + y) = of G f i x + of G f i y := is_add_group_hom.add _ _ _ +@[simp] lemma of_neg (i x) : of G f i (-x) = -of G f i x := is_add_group_hom.neg _ _ +@[simp] lemma of_sub (i x y) : of G f i (x - y) = of G f i x - of G f i y := is_add_group_hom.sub _ _ _ + +@[elab_as_eliminator] +protected theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit G f) + (ih : ∀ i x, C (of G f i x)) : C z := +module.direct_limit.induction_on z ih + +theorem of.zero_exact (i x) (h : of G f i x = 0) : ∃ j hij, f i j hij x = 0 := +module.direct_limit.of.zero_exact h + variables (P : Type u₁) [add_comm_group P] variables (g : Π i, G i → P) [Π i, is_add_group_hom (g i)] variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) -set_option class.instance_max_depth 51 +variables (G f) def rec : direct_limit G f → P := module.direct_limit.rec ℤ ι G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) (λ i, is_add_group_hom.to_linear_map $ g i) Hg +variables {G f} -end direct_limit +instance rec.is_add_group_hom : is_add_group_hom (rec G f P g Hg) := +linear_map.is_add_group_hom _ -end add_comm_group +@[simp] lemma rec_of (i x) : rec G f P g Hg (of G f i x) = g i x := +module.direct_limit.rec_of _ _ _ +@[simp] lemma rec_zero : rec G f P g Hg 0 = 0 := is_add_group_hom.zero _ +@[simp] lemma rec_add (x y) : rec G f P g Hg (x + y) = rec G f P g Hg x + rec G f P g Hg y := is_add_group_hom.add _ _ _ +@[simp] lemma rec_neg (x) : rec G f P g Hg (-x) = -rec G f P g Hg x := is_add_group_hom.neg _ _ +@[simp] lemma rec_sub (x y) : rec G f P g Hg (x - y) = rec G f P g Hg x - rec G f P g Hg y := is_add_group_hom.sub _ _ _ -namespace ring +lemma rec_unique (F : direct_limit G f → P) [is_add_group_hom F] (x) : + F x = rec G f P (λ i x, F $ of G f i x) (λ i j hij x, by rw of_f) x := +direct_limit.induction_on x $ λ i x, by rw rec_of -variables [Π i, ring (G i)] -variables (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_ring_hom (f i j hij)] -variables [add_comm_group.directed_system G f] +end direct_limit -open free_comm_ring +end add_comm_group -def thing1 : set (free_comm_ring Σ i, G i) := -⋃ i j H, set.range $ λ x, of ⟨j, f i j H x⟩ - of ⟨i, x⟩ -def thing2 : set (free_comm_ring Σ i, G i) := -set.range $ λ i, of ⟨i, 1⟩ - 1 +namespace ring -def thing3 : set (free_comm_ring Σ i, G i) := -⋃ i x, set.range $ λ y, of ⟨i, x + y⟩ - (of ⟨i, x⟩ + of ⟨i, y⟩) +variables [Π i, comm_ring (G i)] +variables (f : Π i j, i ≤ j → G i → G j) +variables [Π i j hij, is_ring_hom (f i j hij)] +variables [directed_system G f] -def thing4 : set (free_comm_ring Σ i, G i) := -⋃ i x, set.range $ λ y, of ⟨i, x * y⟩ - (of ⟨i, x⟩ * of ⟨i, y⟩) +open free_comm_ring def direct_limit : Type (max v w) := -(ideal.span (thing1 G f ∪ thing2 G ∪ thing3 G ∪ thing4 G)).quotient +(ideal.span { a | (∃ i j H x, of (⟨j, f i j H x⟩ : Σ i, G i) - of ⟨i, x⟩ = a) ∨ + (∃ i, of (⟨i, 1⟩ : Σ i, G i) - 1 = a) ∨ + (∃ i x y, of (⟨i, x + y⟩ : Σ i, G i) - (of ⟨i, x⟩ + of ⟨i, y⟩) = a) ∨ + (∃ i x y, of (⟨i, x * y⟩ : Σ i, G i) - (of ⟨i, x⟩ * of ⟨i, y⟩) = a) }).quotient namespace direct_limit @@ -317,16 +259,15 @@ comm_ring.to_ring _ def of (i) (x : G i) : direct_limit G f := ideal.quotient.mk _ $ of ⟨i, x⟩ +variables {G f} + instance of.is_ring_hom (i) : is_ring_hom (of G f i) := -{ map_one := ideal.quotient.eq.2 $ subset_span $ or.inl $ or.inl $ or.inr $ set.mem_range_self i, - map_mul := λ x y, ideal.quotient.eq.2 $ subset_span $ or.inr $ set.mem_Union.2 ⟨i, - set.mem_Union.2 ⟨x, set.mem_range_self y⟩⟩, - map_add := λ x y, ideal.quotient.eq.2 $ subset_span $ or.inl $ or.inr $ set.mem_Union.2 ⟨i, - set.mem_Union.2 ⟨x, set.mem_range_self y⟩⟩ } +{ map_one := ideal.quotient.eq.2 $ subset_span $ or.inr $ or.inl ⟨i, rfl⟩, + map_mul := λ x y, ideal.quotient.eq.2 $ subset_span $ or.inr $ or.inr $ or.inr ⟨i, x, y, rfl⟩, + map_add := λ x y, ideal.quotient.eq.2 $ subset_span $ or.inr $ or.inr $ or.inl ⟨i, x, y, rfl⟩ } -@[simp] lemma of_f {i j} (hij : i ≤ j) (x : G i) : of G f j (f i j hij x) = of G f i x := -ideal.quotient.eq.2 $ subset_span $ or.inl $ or.inl $ or.inl $ set.mem_Union.2 ⟨i, set.mem_Union.2 -⟨j, set.mem_Union.2 ⟨hij, set.mem_range_self _⟩⟩⟩ +@[simp] lemma of_f {i j} (hij) (x) : of G f j (f i j hij x) = of G f i x := +ideal.quotient.eq.2 $ subset_span $ or.inl ⟨i, j, hij, x, rfl⟩ @[simp] lemma of_zero (i) : of G f i 0 = 0 := is_ring_hom.map_zero _ @[simp] lemma of_one (i) : of G f i 1 = 1 := is_ring_hom.map_one _ @@ -339,9 +280,9 @@ ideal.quotient.eq.2 $ subset_span $ or.inl $ or.inl $ or.inl $ set.mem_Union.2 theorem exists_of (z : direct_limit G f) : ∃ i x, of G f i x = z := nonempty.elim (by apply_instance) $ assume ind : ι, quotient.induction_on' z $ λ x, free_abelian_group.induction_on x - ⟨ind, 0, of_zero G f ind⟩ + ⟨ind, 0, of_zero ind⟩ (λ s, multiset.induction_on s - ⟨ind, 1, of_one G f ind⟩ + ⟨ind, 1, of_one ind⟩ (λ a s ih, let ⟨i, x⟩ := a, ⟨j, y, hs⟩ := ih, ⟨k, hik, hjk⟩ := directed_order.directed i j in ⟨k, f i k hik x * f j k hjk y, by rw [of_mul, of_f, of_f, hs]; refl⟩)) (λ s ⟨i, x, ih⟩, ⟨i, -x, by rw [of_neg, ih]; refl⟩) @@ -350,54 +291,11 @@ quotient.induction_on' z $ λ x, free_abelian_group.induction_on x @[elab_as_eliminator] theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit G f) (ih : ∀ i x, C (of G f i x)) : C z := -let ⟨i, x, hx⟩ := exists_of G f z in hx ▸ ih i x - -variables (P : Type u₁) [comm_ring P] -variables (g : Π i, G i → P) [Π i, is_ring_hom (g i)] -variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) -include Hg - -open free_comm_ring - -def rec : direct_limit G f → P := -ideal.quotient.lift _ (free_comm_ring.lift $ λ x, g x.1 x.2) begin - suffices : ideal.span (thing1 G f ∪ thing2 G ∪ thing3 G ∪ thing4 G) ≤ - ideal.comap (free_comm_ring.lift (λ (x : Σ (i : ι), G i), g (x.fst) (x.snd))) ⊥, - { intros x hx, exact (mem_bot P).1 (this hx) }, - rw ideal.span_le, intros x hx, - simp only [thing1, thing2, thing3, thing4, set.mem_union, set.mem_Union, set.mem_range] at hx, - rw [mem_coe, ideal.mem_comap, mem_bot], - rcases hx with ⟨⟨i, j, hij, x, rfl⟩ | hx⟩, - { simp only [lift_sub, lift_of, Hg, sub_self] }, - { rcases hx with ⟨i, rfl⟩, simp only [lift_sub, lift_of, lift_one, is_ring_hom.map_one (g i), sub_self] }, - { rcases hx with ⟨i, x, y, rfl⟩, simp only [lift_sub, lift_of, lift_add, is_ring_hom.map_add (g i), sub_self] }, - { rcases hx with ⟨i, x, y, rfl⟩, simp only [lift_sub, lift_of, lift_mul, is_ring_hom.map_mul (g i), sub_self] } -end - -omit Hg - -instance rec.is_ring_hom : is_ring_hom (rec G f P g Hg) := -⟨free_comm_ring.lift_one _, -λ x y, quotient.induction_on₂' x y $ λ p q, free_comm_ring.lift_mul _ _ _, -λ x y, quotient.induction_on₂' x y $ λ p q, free_comm_ring.lift_add _ _ _⟩ - -end direct_limit - -end ring - - -namespace ring - -namespace direct_limit - -variables [Π i, comm_ring (G i)] -variables (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_ring_hom (f i j hij)] -variables [add_comm_group.directed_system G f] - -open free_comm_ring +let ⟨i, x, hx⟩ := exists_of z in hx ▸ ih i x section of_zero_exact_aux attribute [instance, priority 0] classical.dec +variables (G f) lemma of.zero_exact_aux2 {x : free_comm_ring Σ i, G i} {s t} (hxs : is_supported x s) {j k} (hj : ∀ z : Σ i, G i, z ∈ s → z.1 ≤ j) (hk : ∀ z : Σ i, G i, z ∈ t → z.1 ≤ k) (hjk : j ≤ k) (hst : s ⊆ t) : @@ -412,31 +310,29 @@ begin { rintros _ ⟨p, hps, rfl⟩ n ih, rw [restriction_mul, lift_mul, is_ring_hom.map_mul (f j k hjk), ih, restriction_mul, lift_mul, restriction_of, dif_pos hps, lift_pure, restriction_of, dif_pos (hst hps), lift_pure], - dsimp only, rw add_comm_group.directed_system.Hcomp f, refl }, + dsimp only, rw directed_system.Hcomp f, refl }, { rintros x y ihx ihy, rw [restriction_add, lift_add, is_ring_hom.map_add (f j k hjk), ihx, ihy, restriction_add, lift_add] } end +variables {G f} -lemma of.zero_exact_aux {x : free_comm_ring Σ i, G i} (H : x ∈ ideal.span (thing1 G f ∪ thing2 G ∪ thing3 G ∪ thing4 G)) : +lemma of.zero_exact_aux {x : free_comm_ring Σ i, G i} (H : ideal.quotient.mk _ x = (0 : direct_limit G f)) : ∃ j s, ∃ H : (∀ k : Σ i, G i, k ∈ s → k.1 ≤ j), is_supported x s ∧ lift (λ ix : s, f ix.1.1 j (H ix ix.2) ix.1.2) (restriction s x) = (0 : G j) := begin - refine span_induction H _ _ _ _, - { rintros x (⟨⟨hx | hx⟩ | hx⟩ | hx), - { simp only [thing1, set.mem_Union, set.mem_range] at hx, rcases hx with ⟨i, j, hij, x, rfl⟩, - refine ⟨j, {⟨i, x⟩, ⟨j, f i j hij x⟩}, _, + refine span_induction (ideal.quotient.eq_zero_iff_mem.1 H) _ _ _ _, + { rintros x (⟨i, j, hij, x, rfl⟩ | ⟨i, rfl⟩ | ⟨i, x, y, rfl⟩ | ⟨i, x, y, rfl⟩), + { refine ⟨j, {⟨i, x⟩, ⟨j, f i j hij x⟩}, _, is_supported_sub (is_supported_pure.2 $ or.inl rfl) (is_supported_pure.2 $ or.inr $ or.inl rfl), _⟩, { rintros k (rfl | ⟨rfl | h⟩), refl, exact hij, cases h }, { rw [restriction_sub, lift_sub, restriction_of, dif_pos, restriction_of, dif_pos, lift_pure, lift_pure], - dsimp only, rw add_comm_group.directed_system.Hcomp f, exact sub_self _, + dsimp only, rw directed_system.Hcomp f, exact sub_self _, { left, refl }, { right, left, refl }, } }, - { simp only [thing2, set.mem_range] at hx, rcases hx with ⟨i, rfl⟩, - refine ⟨i, {⟨i, 1⟩}, _, is_supported_sub (is_supported_pure.2 $ or.inl rfl) is_supported_one, _⟩, + { refine ⟨i, {⟨i, 1⟩}, _, is_supported_sub (is_supported_pure.2 $ or.inl rfl) is_supported_one, _⟩, { rintros k (rfl | h), refl, cases h }, { rw [restriction_sub, lift_sub, restriction_of, dif_pos, restriction_one, lift_pure, lift_one], dsimp only, rw [is_ring_hom.map_one (f i i _), sub_self], exact _inst_7 i i _, { left, refl } } }, - { simp only [thing3, set.mem_Union, set.mem_range] at hx, rcases hx with ⟨i, x, y, rfl⟩, - refine ⟨i, {⟨i, x+y⟩, ⟨i, x⟩, ⟨i, y⟩}, _, + { refine ⟨i, {⟨i, x+y⟩, ⟨i, x⟩, ⟨i, y⟩}, _, is_supported_sub (is_supported_pure.2 $ or.inr $ or.inr $ or.inl rfl) (is_supported_add (is_supported_pure.2 $ or.inr $ or.inl rfl) (is_supported_pure.2 $ or.inl rfl)), _⟩, { rintros k (rfl | ⟨rfl | ⟨rfl | hk⟩⟩), refl, refl, refl, cases hk }, @@ -444,8 +340,7 @@ begin dif_pos, dif_pos, dif_pos, lift_sub, lift_add, lift_pure, lift_pure, lift_pure], dsimp only, rw is_ring_hom.map_add (f i i _), exact sub_self _, { right, right, left, refl }, { apply_instance }, { left, refl }, { right, left, refl } } }, - { simp only [thing4, set.mem_Union, set.mem_range] at hx, rcases hx with ⟨i, x, y, rfl⟩, - refine ⟨i, {⟨i, x*y⟩, ⟨i, x⟩, ⟨i, y⟩}, _, + { refine ⟨i, {⟨i, x*y⟩, ⟨i, x⟩, ⟨i, y⟩}, _, is_supported_sub (is_supported_pure.2 $ or.inr $ or.inr $ or.inl rfl) (is_supported_mul (is_supported_pure.2 $ or.inr $ or.inl rfl) (is_supported_pure.2 $ or.inl rfl)), _⟩, { rintros k (rfl | ⟨rfl | ⟨rfl | hk⟩⟩), refl, refl, refl, cases hk }, @@ -467,21 +362,21 @@ begin ← of.zero_exact_aux2 G f hyt hj this hjk (set.subset_union_right s t), ihs, is_ring_hom.map_zero (f i k hik), iht, is_ring_hom.map_zero (f j k hjk), zero_add] } }, { rintros x y ⟨j, t, hj, hyt, iht⟩, rw smul_eq_mul, - rcases exists_finset x with ⟨s, hfs, hxs⟩, - rcases (set.finite_image sigma.fst hfs).exists_le with ⟨i, hi⟩, + rcases exists_finset_support x with ⟨s, hxs⟩, + rcases (s.image sigma.fst).exists_le with ⟨i, hi⟩, rcases directed_order.directed i j with ⟨k, hik, hjk⟩, - have : ∀ z : Σ i, G i, z ∈ s ∪ t → z.1 ≤ k, - { rintros z (hz | hz), exact le_trans (hi z.1 ⟨z, hz, rfl⟩) hik, exact le_trans (hj z hz) hjk }, - refine ⟨k, s ∪ t, this, is_supported_mul (is_supported_upwards hxs $ set.subset_union_left s t) - (is_supported_upwards hyt $ set.subset_union_right s t), _⟩, + have : ∀ z : Σ i, G i, z ∈ ↑s ∪ t → z.1 ≤ k, + { rintros z (hz | hz), exact le_trans (hi z.1 $ finset.mem_image.2 ⟨z, hz, rfl⟩) hik, exact le_trans (hj z hz) hjk }, + refine ⟨k, ↑s ∪ t, this, is_supported_mul (is_supported_upwards hxs $ set.subset_union_left ↑s t) + (is_supported_upwards hyt $ set.subset_union_right ↑s t), _⟩, rw [restriction_mul, lift_mul, - ← of.zero_exact_aux2 G f hyt hj this hjk (set.subset_union_right s t), + ← of.zero_exact_aux2 G f hyt hj this hjk (set.subset_union_right ↑s t), iht, is_ring_hom.map_zero (f j k hjk), mul_zero] } end end of_zero_exact_aux lemma of.zero_exact {i x} (hix : of G f i x = 0) : ∃ j, ∃ hij : i ≤ j, f i j hij x = 0 := -let ⟨j, s, H, hxs, hx⟩ := of.zero_exact_aux G f (ideal.quotient.eq_zero_iff_mem.1 hix) in +let ⟨j, s, H, hxs, hx⟩ := of.zero_exact_aux hix in have hixs : (⟨i, x⟩ : Σ i, G i) ∈ s, from is_supported_pure.1 hxs, ⟨j, H ⟨i, x⟩ hixs, by rw [restriction_of, dif_pos hixs, lift_pure] at hx; exact hx⟩ @@ -491,10 +386,50 @@ begin suffices : ∀ x, of G f i x = 0 → x = 0, { intros x y hxy, rw ← sub_eq_zero_iff_eq, apply this, rw [is_ring_hom.map_sub (of G f i), hxy, sub_self] }, - intros x hx, rcases of.zero_exact G f hx with ⟨j, hij, hfx⟩, + intros x hx, rcases of.zero_exact hx with ⟨j, hij, hfx⟩, apply hf i j hij, rw [hfx, is_ring_hom.map_zero (f i j hij)] end +variables (P : Type u₁) [comm_ring P] +variables (g : Π i, G i → P) [Π i, is_ring_hom (g i)] +variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) +include Hg + +open free_comm_ring + +variables (G f) +def rec : direct_limit G f → P := +ideal.quotient.lift _ (free_comm_ring.lift $ λ x, g x.1 x.2) begin + suffices : ideal.span _ ≤ + ideal.comap (free_comm_ring.lift (λ (x : Σ (i : ι), G i), g (x.fst) (x.snd))) ⊥, + { intros x hx, exact (mem_bot P).1 (this hx) }, + rw ideal.span_le, intros x hx, + rw [mem_coe, ideal.mem_comap, mem_bot], + rcases hx with ⟨i, j, hij, x, rfl⟩ | ⟨i, rfl⟩ | ⟨i, x, y, rfl⟩ | ⟨i, x, y, rfl⟩; + simp only [lift_sub, lift_of, Hg, lift_one, lift_add, lift_mul, + is_ring_hom.map_one (g i), is_ring_hom.map_add (g i), is_ring_hom.map_mul (g i), sub_self] +end +variables {G f} +omit Hg + +instance rec.is_ring_hom : is_ring_hom (rec G f P g Hg) := +⟨free_comm_ring.lift_one _, +λ x y, quotient.induction_on₂' x y $ λ p q, free_comm_ring.lift_mul _ _ _, +λ x y, quotient.induction_on₂' x y $ λ p q, free_comm_ring.lift_add _ _ _⟩ + +@[simp] lemma rec_of (i x) : rec G f P g Hg (of G f i x) = g i x := free_comm_ring.lift_of _ _ +@[simp] lemma rec_zero : rec G f P g Hg 0 = 0 := is_ring_hom.map_zero _ +@[simp] lemma rec_one : rec G f P g Hg 1 = 1 := is_ring_hom.map_one _ +@[simp] lemma rec_add (x y) : rec G f P g Hg (x + y) = rec G f P g Hg x + rec G f P g Hg y := is_ring_hom.map_add _ +@[simp] lemma rec_neg (x) : rec G f P g Hg (-x) = -rec G f P g Hg x := is_ring_hom.map_neg _ +@[simp] lemma rec_sub (x y) : rec G f P g Hg (x - y) = rec G f P g Hg x - rec G f P g Hg y := is_ring_hom.map_sub _ +@[simp] lemma rec_mul (x y) : rec G f P g Hg (x * y) = rec G f P g Hg x * rec G f P g Hg y := is_ring_hom.map_mul _ +@[simp] lemma rec_pow (x) (n : ℕ) : rec G f P g Hg (x ^ n) = rec G f P g Hg x ^ n := is_semiring_hom.map_pow _ _ _ + +theorem rec_unique (F : direct_limit G f → P) [is_ring_hom F] (x) : + F x = rec G f P (λ i x, F $ of G f i x) (λ i j hij x, by rw [of_f]) x := +direct_limit.induction_on x $ λ i x, by rw rec_of + end direct_limit end ring @@ -504,51 +439,44 @@ namespace field variables [Π i, field (G i)] variables (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_field_hom (f i j hij)] -variables [add_comm_group.directed_system G f] +variables [directed_system G f] + +namespace direct_limit -instance direct_limit.nonzero_comm_ring : nonzero_comm_ring (ring.direct_limit G f) := +instance nonzero_comm_ring : nonzero_comm_ring (ring.direct_limit G f) := { zero_ne_one := nonempty.elim (by apply_instance) $ assume i : ι, begin change (0 : ring.direct_limit G f) ≠ 1, rw ← ring.direct_limit.of_one, - intros H, rcases ring.direct_limit.of.zero_exact G f H.symm with ⟨j, hij, hf⟩, + intros H, rcases ring.direct_limit.of.zero_exact H.symm with ⟨j, hij, hf⟩, rw is_ring_hom.map_one (f i j hij) at hf, exact one_ne_zero hf end, .. ring.direct_limit.comm_ring G f } -instance direct_limit.is_division_ring : is_division_ring (ring.direct_limit G f) := -{ exists_inv := λ p, ring.direct_limit.induction_on G f p $ λ i x H, - ⟨ring.direct_limit.of G f i (x⁻¹), by erw [← ring.direct_limit.of_mul, - mul_inv_cancel (assume h : x = 0, H $ by rw [h, ring.direct_limit.of_zero]), - ring.direct_limit.of_one]⟩ } - -def direct_limit : Type (max v w) := -to_field (ring.direct_limit G f) - -namespace direct_limit - -set_option class.instance_max_depth 10 -instance : field (direct_limit G f) := -@to_field.field (ring.direct_limit G f) (field.direct_limit.nonzero_comm_ring G f) (field.direct_limit.is_division_ring G f) +theorem exists_inv {p : ring.direct_limit G f} : p ≠ 0 → ∃ y, p * y = 1 := +ring.direct_limit.induction_on p $ λ i x H, +⟨ring.direct_limit.of G f i (x⁻¹), by erw [← ring.direct_limit.of_mul, + mul_inv_cancel (assume h : x = 0, H $ by rw [h, ring.direct_limit.of_zero]), + ring.direct_limit.of_one]⟩ -def of (i) (x : G i) : direct_limit G f := -to_field.mk $ ring.direct_limit.of G f i x +section +local attribute [instance, priority 0] classical.dec -set_option class.instance_max_depth 20 -instance of.is_ring_hom (i) : is_ring_hom (of G f i) := -@is_ring_hom.comp _ _ _ _ (ring.direct_limit.of G f i) (ring.direct_limit.of.is_ring_hom G f i) - _ _ (to_field.mk) (to_field.mk.is_ring_hom _) +noncomputable def inv (p : ring.direct_limit G f) : ring.direct_limit G f := +if H : p = 0 then 0 else classical.some (direct_limit.exists_inv G f H) -theorem of_f {i j} (hij : i ≤ j) (x : G i) : of G f j (f i j hij x) = of G f i x := -congr_arg to_field.mk $ ring.direct_limit.of_f G f hij x +protected theorem mul_inv_cancel {p : ring.direct_limit G f} (hp : p ≠ 0) : p * inv G f p = 1 := +by rw [inv, dif_neg hp, classical.some_spec (direct_limit.exists_inv G f hp)] -variables (P : Type u₁) [field' P] -variables (g : Π i, G i → P) [Π i, is_field_hom (g i)] -variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) -include Hg +protected theorem inv_mul_cancel {p : ring.direct_limit G f} (hp : p ≠ 0) : inv G f p * p = 1 := +by rw [_root_.mul_comm, direct_limit.mul_inv_cancel G f hp] -def rec : direct_limit G f → P := -to_field.eval $ ring.direct_limit.rec G f P g Hg +protected noncomputable def field : field (ring.direct_limit G f) := +{ inv := inv G f, + mul_inv_cancel := λ p, direct_limit.mul_inv_cancel G f, + inv_mul_cancel := λ p, direct_limit.inv_mul_cancel G f, + .. direct_limit.nonzero_comm_ring G f } +end end direct_limit diff --git a/src/algebra/module.lean b/src/algebra/module.lean index d79e7199bada3..2c7da8cb0e146 100644 --- a/src/algebra/module.lean +++ b/src/algebra/module.lean @@ -108,6 +108,13 @@ instance semiring.to_semimodule [r : semiring α] : semimodule α α := instance ring.to_module [r : ring α] : module α α := { ..semiring.to_semimodule } +instance add_comm_group.module [add_comm_group α] : module ℤ α := module.of_core +{ smul := gsmul, + mul_smul := λ r s x, gsmul_mul x r s, + smul_add := λ r x y, gsmul_add x y r, + add_smul := λ r s x, add_gsmul x r s, + one_smul := one_gsmul } + class is_linear_map (α : Type u) {β : Type v} {γ : Type w} [ring α] [add_comm_group β] [add_comm_group γ] [module α β] [module α γ] (f : β → γ) : Prop := @@ -169,6 +176,14 @@ def id : β →ₗ[α] β := ⟨id, by simp, by simp⟩ end linear_map +def is_add_group_hom.to_linear_map [add_comm_group α] [add_comm_group β] + (f : α → β) [is_add_group_hom f] : α →ₗ[ℤ] β := +{ to_fun := f, + add := is_add_group_hom.add f, + smul := λ i x, int.induction_on i (by rw [zero_smul, zero_smul, is_add_group_hom.zero f]) + (λ i ih, by rw [add_smul, add_smul, is_add_group_hom.add f, ih, one_smul, one_smul]) + (λ i ih, by rw [sub_smul, sub_smul, is_add_group_hom.sub f, ih, one_smul, one_smul]) } + namespace is_linear_map variables [ring α] [add_comm_group β] [add_comm_group γ] variables [module α β] [module α γ] diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 00d6f8122e43e..ac17fc386ef09 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -586,6 +586,13 @@ span_eq_bot.trans $ by simp span_eq_of_le _ (image_subset _ subset_span) $ map_le_iff_le_comap.2 $ span_le.2 $ image_subset_iff.1 subset_span +lemma span_preimage_le_comap_span (f : β →ₗ[α] γ) (s : set γ) : span α (f ⁻¹' s) ≤ (span α s).comap f := +λ x h, span_induction h + (by simp only [set.preimage, set.mem_set_of_eq, mem_comap]; exact λ x h, subset_span h) + (zero_mem ((span α s).comap f)) + (λ _ _ hx hy, add_mem ((span α s).comap f) hx hy) + (λ _ _ h, smul_mem ((span α s).comap f) _ h) + def prod : submodule α (β × γ) := { carrier := set.prod p q, zero := ⟨zero_mem _, zero_mem _⟩, diff --git a/src/order/basic.lean b/src/order/basic.lean index 4c3ee3393fa1b..e568045266d74 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -464,4 +464,7 @@ theorem directed_mono {s : α → α → Prop} {ι} (f : ι → α) (H : ∀ a b, r a b → s a b) (h : directed r f) : directed s f := λ a b, let ⟨c, h₁, h₂⟩ := h a b in ⟨c, H _ _ h₁, H _ _ h₂⟩ +class directed_order (α : Type u) extends partial_order α := +(directed : ∀ i j : α, ∃ k, i ≤ k ∧ j ≤ k) + end diff --git a/src/ring_theory/free_comm_ring.lean b/src/ring_theory/free_comm_ring.lean index 6a418e15cea5a..1ac6818e75a46 100644 --- a/src/ring_theory/free_comm_ring.lean +++ b/src/ring_theory/free_comm_ring.lean @@ -1,3 +1,9 @@ +/- +Copyright (c) 2019 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau +-/ + import group_theory.free_abelian_group ring_theory.subring data.set.finite data.real.irrational universes u v @@ -247,7 +253,7 @@ begin { intros x y ihx ihy, rw [restriction_add, map_add, ihx, ihy] } end -theorem exists_finset (x : free_comm_ring α) : ∃ s : set α, set.finite s ∧ is_supported x s := +theorem exists_finite_support (x : free_comm_ring α) : ∃ s : set α, set.finite s ∧ is_supported x s := free_comm_ring.induction_on x ⟨∅, set.finite_empty, is_supported_neg is_supported_one⟩ (λ p, ⟨{p}, set.finite_singleton p, is_supported_pure.2 $ finset.mem_singleton_self _⟩) @@ -258,4 +264,7 @@ free_comm_ring.induction_on x (is_supported_upwards hxs $ set.subset_union_left s t) (is_supported_upwards hxt $ set.subset_union_right s t)⟩) +theorem exists_finset_support (x : free_comm_ring α) : ∃ s : finset α, is_supported x ↑s := +let ⟨s, hfs, hxs⟩ := exists_finite_support x in ⟨hfs.to_finset, by rwa finset.coe_to_finset⟩ + end free_comm_ring diff --git a/src/ring_theory/to_field.lean b/src/ring_theory/to_field.lean deleted file mode 100644 index 246729782a7d5..0000000000000 --- a/src/ring_theory/to_field.lean +++ /dev/null @@ -1,231 +0,0 @@ -import algebra.ring - -universes u v - -variables (α : Type u) - -class is_division_ring [ring α] : Prop := -(exists_inv : ∀ {x : α}, x ≠ 0 → ∃ y, x * y = 1) - -section -set_option old_structure_cmd true -class field' (α : Type u) extends division_ring α, comm_ring α := -(inv_zero : (0:α)⁻¹ = 0) -end - -theorem inv_zero' [field' α] : (0:α)⁻¹ = 0 := -field'.inv_zero α - -theorem mul_eq_one_comm {α} [ring α] [is_division_ring α] {x y : α} (hxy : x * y = 1) : y * x = 1 := -classical.by_cases - (assume hy0 : y = 0, by rw [← hxy, hy0, zero_mul, mul_zero]) - (assume hy0 : y ≠ 0, let ⟨z, hz⟩ := is_division_ring.exists_inv hy0 in - by rw [← mul_one x, ← hz, ← mul_assoc x, hxy, one_mul]) - -inductive pre_to_field : Type u -| of : α → pre_to_field -| add : pre_to_field → pre_to_field → pre_to_field -| mul : pre_to_field → pre_to_field → pre_to_field -| inv : pre_to_field → pre_to_field - -namespace pre_to_field - -variables {α} - -def rel [semiring α] : (pre_to_field α) → α → Prop -| (of p) z := p = z -| (add p q) z := ∃ x y, rel p x ∧ rel q y ∧ x + y = z -| (mul p q) z := ∃ x y, rel p x ∧ rel q y ∧ x * y = z -| (inv p) z := rel p 0 ∧ z = 0 ∨ ∃ x, rel p x ∧ x * z = 1 - -def neg [has_neg α] : pre_to_field α → pre_to_field α -| (of p) := of (-p) -| (add x y) := add (neg x) (neg y) -| (mul x y) := mul (neg x) y -| (inv x) := inv (neg x) - -theorem rel_neg [ring α] : Π {p} {z : α}, rel p z → rel (neg p) (-z) -| (of p) z hpz := congr_arg has_neg.neg hpz -| (add p q) z ⟨x, y, hpx, hqy, hxyz⟩ := ⟨-x, -y, rel_neg hpx, rel_neg hqy, by rw [← hxyz, neg_add]⟩ -| (mul p q) z ⟨x, y, hpx, hqy, hxyz⟩ := ⟨-x, y, rel_neg hpx, hqy, by rw [← hxyz, neg_mul_eq_neg_mul]⟩ -| (inv p) z (or.inl ⟨hp0, hz0⟩) := or.inl ⟨@neg_zero α _ ▸ rel_neg hp0, neg_eq_zero.2 hz0⟩ -| (inv p) z (or.inr ⟨x, hpx, hxz⟩) := or.inr ⟨-x, rel_neg hpx, by rw [neg_mul_neg, hxz]⟩ - -variables [ring α] [is_division_ring α] - -lemma unique_rel (p : pre_to_field α) : - ∃!x, rel p x := -pre_to_field.rec_on p - (λ x, ⟨x, rfl, λ y hxy, hxy.symm⟩) - (λ p q ⟨x, hpx, hpx2⟩ ⟨y, hqy, hqy2⟩, ⟨x+y, ⟨x, y, hpx, hqy, rfl⟩, - λ z ⟨x', y', hpx', hqy', hxyz⟩, hpx2 x' hpx' ▸ hqy2 y' hqy' ▸ hxyz.symm⟩) - (λ p q ⟨x, hpx, hpx2⟩ ⟨y, hqy, hqy2⟩, ⟨x*y, ⟨x, y, hpx, hqy, rfl⟩, - λ z ⟨x', y', hpx', hqy', hxyz⟩, hpx2 x' hpx' ▸ hqy2 y' hqy' ▸ hxyz.symm⟩) - (λ p ⟨x, hpx, hpx2⟩, classical.by_cases - (assume hx0 : x = 0, ⟨0, or.inl ⟨hx0 ▸ hpx, rfl⟩, λ y hy, - or.cases_on hy and.right $ λ ⟨z, hpz, hzy⟩, - by rw [← one_mul y, ← hzy, hpx2 z hpz, hx0, zero_mul, zero_mul]⟩) - (assume hx0 : x ≠ 0, let ⟨y, hxy⟩ := is_division_ring.exists_inv hx0 in - ⟨y, or.inr ⟨x, hpx, hxy⟩, λ z hz, or.cases_on hz - (λ ⟨hp0, hz0⟩, by rw [hz0, ← one_mul y, ← hxy, ← hpx2 0 hp0, zero_mul, zero_mul]) - (λ ⟨s, hps, hsz⟩, by rw [← mul_one y, ← hsz, ← mul_assoc, hpx2 s hps, mul_eq_one_comm hxy, one_mul])⟩)) - -variables {β : Type v} - -def eval [division_ring β] (f : α → β) : pre_to_field α → β -| (of p) := f p -| (add x y) := eval x + eval y -| (mul x y) := eval x * eval y -| (inv p) := (eval p)⁻¹ - -theorem eval_of_rel [field' β] (f : α → β) [is_ring_hom f] : - Π {p} {x : α} (hpx : pre_to_field.rel p x), eval f p = f x -| (of p) z hpz := congr_arg f hpz -| (add p q) z ⟨x, y, hpx, hqy, hxyz⟩ := by rw [eval, eval_of_rel hpx, eval_of_rel hqy, ← hxyz, is_ring_hom.map_add f] -| (mul p q) z ⟨x, y, hpx, hqy, hxyz⟩ := by rw [eval, eval_of_rel hpx, eval_of_rel hqy, ← hxyz, is_ring_hom.map_mul f] -| (inv p) z (or.inl ⟨hp0, hz0⟩) := by rw [eval, eval_of_rel hp0, hz0, is_ring_hom.map_zero f, inv_zero'] -| (inv p) z (or.inr ⟨x, hpx, hxz⟩) := by rw [eval, eval_of_rel hpx, ← mul_one (f x)⁻¹, ← is_ring_hom.map_one f, - ← hxz, is_ring_hom.map_mul f, ← mul_assoc, - inv_mul_cancel (assume hfx : f x = 0, @zero_ne_one β _ $ - by rw [← is_ring_hom.map_one f, ← hxz, is_ring_hom.map_mul f, hfx, zero_mul]), - one_mul] - -variables (α) - -instance : setoid (pre_to_field α) := -{ r := λ p q, ∃ x, rel p x ∧ rel q x, - iseqv := ⟨λ p, let ⟨x, hx⟩ := unique_rel p in ⟨x, hx.1, hx.1⟩, - λ p q ⟨x, hpx, hqx⟩, ⟨x, hqx, hpx⟩, - λ p q r ⟨x, hpx, hqx⟩ ⟨y, hqy, hry⟩, ⟨x, hpx, unique_of_exists_unique (unique_rel q) hqy hqx ▸ hry⟩⟩ } - -end pre_to_field - - -def to_field [ring α] [is_division_ring α] : Type u := -quotient (pre_to_field.setoid α) - -namespace to_field - -variables {α} [ring α] [is_division_ring α] - -def mk (x : α) : to_field α := -⟦pre_to_field.of x⟧ - -variables (α) - -theorem mk.bijective : function.bijective (@mk α _ _) := -⟨λ x y H, let ⟨z, hxz, hyz⟩ := quotient.exact H in eq.trans hxz hyz.symm, -λ x, quotient.induction_on x $ λ p, let ⟨z, hz⟩ := pre_to_field.unique_rel p in -⟨z, quotient.sound ⟨z, rfl, hz.1⟩⟩⟩ - -instance : has_add (to_field α) := -⟨λ x y, quotient.lift_on₂ x y (λ p q, ⟦pre_to_field.add p q⟧) $ λ p q r s ⟨x, hpx, hrx⟩ ⟨y, hqy, hsy⟩, quotient.sound -⟨x + y, ⟨x, y, hpx, hqy, rfl⟩, ⟨x, y, hrx, hsy, rfl⟩⟩⟩ - -instance : has_neg (to_field α) := -⟨λ x, quotient.lift_on x (λ p, ⟦pre_to_field.neg p⟧) $ λ p q ⟨x, hpx, hqx⟩, quotient.sound -⟨-x, pre_to_field.rel_neg hpx, pre_to_field.rel_neg hqx⟩⟩ - -instance : has_mul (to_field α) := -⟨λ x y, quotient.lift_on₂ x y (λ p q, ⟦pre_to_field.mul p q⟧) $ λ p q r s ⟨x, hpx, hrx⟩ ⟨y, hqy, hsy⟩, quotient.sound -⟨x * y, ⟨x, y, hpx, hqy, rfl⟩, ⟨x, y, hrx, hsy, rfl⟩⟩⟩ - -instance : has_inv (to_field α) := -⟨λ x, quotient.lift_on x (λ p, ⟦pre_to_field.inv p⟧) $ λ p q ⟨x, hpx, hqx⟩, quotient.sound $ -classical.by_cases - (assume hx0 : x = 0, ⟨0, or.inl ⟨hx0 ▸ hpx, rfl⟩, or.inl ⟨hx0 ▸ hqx, rfl⟩⟩) - (assume hx0 : x ≠ 0, let ⟨y, hy⟩ := is_division_ring.exists_inv hx0 in ⟨y, or.inr ⟨x, hpx, hy⟩, or.inr ⟨x, hqx, hy⟩⟩)⟩ - -variables {α} - -@[elab_as_eliminator] protected lemma induction_on - {C : to_field α → Prop} (x : to_field α) - (ih : ∀ z, C (mk z)) : C x := -let ⟨z, hz⟩ := (mk.bijective α).2 x in hz ▸ ih z - -@[elab_as_eliminator] protected lemma induction_on₂ - {C : to_field α → to_field α → Prop} (x y : to_field α) - (ih : ∀ p q, C (mk p) (mk q)) : C x y := -to_field.induction_on x $ λ p, to_field.induction_on y $ λ q, ih p q - -@[elab_as_eliminator] protected lemma induction_on₃ - {C : to_field α → to_field α → to_field α → Prop} (x y z : to_field α) - (ih : ∀ p q r, C (mk p) (mk q) (mk r)) : C x y z := -to_field.induction_on x $ λ p, to_field.induction_on y $ λ q, to_field.induction_on z $ λ r, ih p q r - -@[simp] lemma mk_add (x y : α) : mk (x + y) = mk x + mk y := -quotient.sound ⟨x + y, rfl, ⟨x, y, rfl, rfl, rfl⟩⟩ - -@[simp] lemma mk_neg (x : α) : mk (-x) = -mk x := rfl - -@[simp] lemma mk_mul (x y : α) : mk (x * y) = mk x * mk y := -quotient.sound ⟨x * y, rfl, ⟨x, y, rfl, rfl, rfl⟩⟩ - -lemma mk_inv {x y : α} (hxy : x * y = 1) : (mk x)⁻¹ = mk y := -quotient.sound ⟨y, or.inr ⟨x, rfl, hxy⟩, rfl⟩ - -variables (α) -instance : ring (to_field α) := -{ zero := mk 0, - one := mk 1, - add_assoc := λ x y z, to_field.induction_on₃ x y z $ λ p q r, - by rw [← mk_add, ← mk_add, ← mk_add, ← mk_add, add_assoc], - zero_add := λ x, to_field.induction_on x $ λ p, - by rw [← mk_add, zero_add], - add_zero := λ x, to_field.induction_on x $ λ p, - by rw [← mk_add, add_zero], - add_left_neg := λ x, to_field.induction_on x $ λ p, - by rw [← mk_neg, ← mk_add, add_left_neg]; refl, - add_comm := λ x y, to_field.induction_on₂ x y $ λ p q, - by rw [← mk_add, ← mk_add, add_comm], - mul_assoc := λ x y z, to_field.induction_on₃ x y z $ λ p q r, - by rw [← mk_mul, ← mk_mul, ← mk_mul, ← mk_mul, mul_assoc], - one_mul := λ x, to_field.induction_on x $ λ p, - by rw [← mk_mul, one_mul], - mul_one := λ x, to_field.induction_on x $ λ p, - by rw [← mk_mul, mul_one], - left_distrib := λ x y z, to_field.induction_on₃ x y z $ λ p q r, - by rw [← mk_add, ← mk_mul, ← mk_mul, ← mk_mul, ← mk_add, mul_add], - right_distrib := λ x y z, to_field.induction_on₃ x y z $ λ p q r, - by rw [← mk_add, ← mk_mul, ← mk_mul, ← mk_mul, ← mk_add, add_mul], - .. to_field.has_add α, - .. to_field.has_neg α, - .. to_field.has_mul α } - -@[simp] lemma mk_zero : mk (0 : α) = 0 := rfl -@[simp] lemma mk_one : mk (1 : α) = 1 := rfl - -instance mk.is_ring_hom : is_ring_hom (@mk α _ _) := -⟨mk_one α, mk_mul, mk_add⟩ - -variables {α} -def eval {β : Type v} [field' β] (f : α → β) [is_ring_hom f] (p : to_field α) : β := -quotient.lift_on p (pre_to_field.eval f) $ λ p q ⟨x, hpx, hqx⟩, -by rw [pre_to_field.eval_of_rel f hpx, pre_to_field.eval_of_rel f hqx] - -instance eval.is_ring_hom {β : Type v} [field' β] (f : α → β) [is_ring_hom f] : is_ring_hom (eval f) := -⟨by convert is_ring_hom.map_one f, -λ p q, quotient.induction_on₂ p q $ λ x y, rfl, -λ p q, quotient.induction_on₂ p q $ λ x y, rfl⟩ - -end to_field - -namespace to_field - -instance [comm_ring α] [is_division_ring α] : comm_ring (to_field α) := -{ mul_comm := λ x y, to_field.induction_on₂ x y $ λ p q, - by rw [← mk_mul, ← mk_mul, mul_comm], - .. to_field.ring α } - -instance [nonzero_comm_ring α] [is_division_ring α] : field (to_field α) := -{ zero_ne_one := λ H, zero_ne_one $ (mk.bijective α).1 H, - mul_inv_cancel := λ x, to_field.induction_on x $ λ p hp0, - let ⟨y, hy⟩ := is_division_ring.exists_inv (mt (congr_arg mk) hp0) in - by rw [mk_inv hy, ← mk_mul, hy, mk_one], - inv_mul_cancel := λ x, to_field.induction_on x $ λ p hp0, - let ⟨y, hy⟩ := is_division_ring.exists_inv (mt (congr_arg mk) hp0) in - by rw [mk_inv hy, ← mk_mul, mul_eq_one_comm hy, mk_one], - .. to_field.comm_ring α, - .. to_field.has_inv α } - -end to_field From f2be1d590ae6dacc0af14579afa2c37cd5cd252c Mon Sep 17 00:00:00 2001 From: kckennylau Date: Sat, 23 Feb 2019 18:58:28 +0000 Subject: [PATCH 16/22] remove ununsed lemmas --- src/group_theory/free_abelian_group.lean | 32 ----------------------- src/linear_algebra/basic.lean | 7 ----- src/linear_algebra/direct_sum_module.lean | 5 ---- 3 files changed, 44 deletions(-) diff --git a/src/group_theory/free_abelian_group.lean b/src/group_theory/free_abelian_group.lean index f620a72f2b0de..bf589c1642475 100644 --- a/src/group_theory/free_abelian_group.lean +++ b/src/group_theory/free_abelian_group.lean @@ -115,36 +115,4 @@ begin ac_refl } end -section count - -variables [decidable_eq α] - -def count (x : α) : free_abelian_group α → ℤ := -lift $ λ y, if x = y then 1 else 0 - -theorem count_of (x y : α) : count x (of y) = if x = y then 1 else 0 := -lift.of _ _ - -theorem count_of_self (x : α) : count x (of x) = 1 := -(count_of x x).trans $ if_pos rfl - -theorem count_zero (x : α) : count x 0 = 0 := -lift.zero _ - -end count - -theorem of_inj (x y : α) (H : of x = of y) : x = y := -begin - classical, replace H := congr_arg (count x) H, - rw [count_of_self x, count_of] at H, split_ifs at H with h, - { exact h }, { exfalso, exact one_ne_zero H } -end - -theorem of_ne_zero (x : α) : of x ≠ 0 := -begin - intros H, classical, replace H := congr_arg (count x) H, - rw [count_of_self x, count_zero] at H, - exact one_ne_zero H -end - end free_abelian_group diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index ac17fc386ef09..00d6f8122e43e 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -586,13 +586,6 @@ span_eq_bot.trans $ by simp span_eq_of_le _ (image_subset _ subset_span) $ map_le_iff_le_comap.2 $ span_le.2 $ image_subset_iff.1 subset_span -lemma span_preimage_le_comap_span (f : β →ₗ[α] γ) (s : set γ) : span α (f ⁻¹' s) ≤ (span α s).comap f := -λ x h, span_induction h - (by simp only [set.preimage, set.mem_set_of_eq, mem_comap]; exact λ x h, subset_span h) - (zero_mem ((span α s).comap f)) - (λ _ _ hx hy, add_mem ((span α s).comap f) hx hy) - (λ _ _ h, smul_mem ((span α s).comap f) _ h) - def prod : submodule α (β × γ) := { carrier := set.prod p q, zero := ⟨zero_mem _, zero_mem _⟩, diff --git a/src/linear_algebra/direct_sum_module.lean b/src/linear_algebra/direct_sum_module.lean index d9afd5044c742..060109f2198e5 100644 --- a/src/linear_algebra/direct_sum_module.lean +++ b/src/linear_algebra/direct_sum_module.lean @@ -39,11 +39,6 @@ theorem mk_smul (s : finset ι) (c : R) (x) : mk β s (c • x) = c • mk β s theorem of_smul (i : ι) (c : R) (x) : of β i (c • x) = c • of β i x := (lof R ι β i).map_smul c x -lemma sum_of [Π i, decidable_pred (eq (0 : β i))] (f : direct_sum ι β) : - f.sum (λ i, lof R ι β i) = f := -by dsimp [of, dfinsupp.lsingle]; unfold_coes; - exact @dfinsupp.sum_single ι β _ _ _ f - variables {γ : Type u₁} [add_comm_group γ] [module R γ] variables (φ : Π i, β i →ₗ[R] γ) From 83542ff5437d607b8e59bfdb5f4b84ff282173a3 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Sun, 24 Feb 2019 07:48:06 +0000 Subject: [PATCH 17/22] clean up --- src/algebra/direct_limit.lean | 93 ++++++++++++++++++----------------- 1 file changed, 47 insertions(+), 46 deletions(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 606aab7b1ee02..44d6ebe157f49 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -21,16 +21,16 @@ variables [directed_order ι] [decidable_eq ι] variables (G : ι → Type w) [Π i, decidable_eq (G i)] class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := -(Hid : ∀ i x h, f i i h x = x) -(Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) +(map_self : ∀ i x h, f i i h x = x) +(map_map : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) namespace module variables [Π i, add_comm_group (G i)] [Π i, module R (G i)] class directed_system (f : Π i j, i ≤ j → G i →ₗ[R] G j) : Prop := -(Hid : ∀ i x h, f i i h x = x) -(Hcomp : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) +(map_self : ∀ i x h, f i i h x = x) +(map_map : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) variables (f : Π i j, i ≤ j → G i →ₗ[R] G j) [directed_system G f] @@ -48,41 +48,41 @@ def of (i) : G i →ₗ[R] direct_limit G f := (mkq _).comp $ direct_sum.lof R ι G i variables {R ι G f} -theorem of_f {i j hij x} : (of R ι G f j (f i j hij x)) = of R ι G f i x := +@[simp] lemma of_f {i j hij x} : (of R ι G f j (f i j hij x)) = of R ι G f i x := eq.symm $ (submodule.quotient.eq _).2 $ subset_span ⟨i, j, hij, x, rfl⟩ -theorem exists_of (z : direct_limit G f) : ∃ i x, z = of R ι G f i x := +theorem exists_of (z : direct_limit G f) : ∃ i x, of R ι G f i x = z := nonempty.elim (by apply_instance) $ assume ind : ι, quotient.induction_on' z $ λ z, direct_sum.induction_on z - ⟨ind, 0, (linear_map.map_zero _).symm⟩ + ⟨ind, 0, linear_map.map_zero _⟩ (λ i x, ⟨i, x, rfl⟩) (λ p q ⟨i, x, ihx⟩ ⟨j, y, ihy⟩, let ⟨k, hik, hjk⟩ := directed_order.directed i j in - ⟨k, f i k hik x + f j k hjk y, by rw [linear_map.map_add, of_f, of_f, ← ihx, ← ihy]; refl⟩) + ⟨k, f i k hik x + f j k hjk y, by rw [linear_map.map_add, of_f, of_f, ihx, ihy]; refl⟩) @[elab_as_eliminator] protected theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit G f) (ih : ∀ i x, C (of R ι G f i x)) : C z := -let ⟨i, x, h⟩ := exists_of z in h.symm ▸ ih i x +let ⟨i, x, h⟩ := exists_of z in h ▸ ih i x variables {P : Type u₁} [add_comm_group P] [module R P] (g : Π i, G i →ₗ[R] P) variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) include Hg variables (R ι G f) -def rec : direct_limit G f →ₗ[R] P := +def lift : direct_limit G f →ₗ[R] P := liftq _ (direct_sum.to_module R ι P g) (span_le.2 $ λ a ⟨i, j, hij, x, hx⟩, by rw [← hx, mem_coe, linear_map.sub_mem_ker_iff, direct_sum.to_module_lof, direct_sum.to_module_lof, Hg]) variables {R ι G f} omit Hg -lemma rec_of {i} (x) : rec R ι G f g Hg (of R ι G f i x) = g i x := +lemma lift_of {i} (x) : lift R ι G f g Hg (of R ι G f i x) = g i x := direct_sum.to_module_lof R _ _ -theorem rec_unique (F : direct_limit G f →ₗ[R] P) (x) : - F x = rec R ι G f (λ i, F.comp $ of R ι G f i) +theorem lift_unique (F : direct_limit G f →ₗ[R] P) (x) : + F x = lift R ι G f (λ i, F.comp $ of R ι G f i) (λ i j hij x, by rw [linear_map.comp_apply, of_f]; refl) x := -direct_limit.induction_on x $ λ i x, by rw rec_of; refl +direct_limit.induction_on x $ λ i x, by rw lift_of; refl section totalize local attribute [instance, priority 0] classical.dec @@ -107,7 +107,7 @@ begin simp only [linear_map.map_sum], refine finset.sum_congr rfl (λ k hk, _), rw direct_sum.single_eq_lof R k (x k), - simp [totalize_apply, hx k hk, le_trans (hx k hk) hij, directed_system.Hcomp f] + simp [totalize_apply, hx k hk, le_trans (hx k hk) hij, directed_system.map_map f] end lemma of.zero_exact_aux {x : direct_sum ι G} (H : submodule.quotient.mk x = (0 : direct_limit G f)) : @@ -125,7 +125,7 @@ span_induction ((quotient.mk_eq_zero _).1 H) split_ifs at hi0 with hi hj hj, { rwa hi at hik }, { rwa hi at hik }, { rwa hj at hjk }, exfalso, apply hi0, rw sub_zero }, simp [linear_map.map_sub, totalize_apply, hik, hjk, - directed_system.Hcomp f, direct_sum.apply_eq_component, + directed_system.map_map f, direct_sum.apply_eq_component, direct_sum.component.of], end⟩) ⟨ind, λ _ h, (finset.not_mem_empty _ h).elim, linear_map.map_zero _⟩ @@ -164,7 +164,7 @@ def direct_limit (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] : Type* := @module.direct_limit ℤ _ ι _ _ _ G _ _ _ (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) - ⟨directed_system.Hid f, directed_system.Hcomp f⟩ + ⟨directed_system.map_self f, directed_system.map_map f⟩ namespace direct_limit @@ -172,7 +172,7 @@ variables (f : Π i j, i ≤ j → G i → G j) variables [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] def directed_system : module.directed_system G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) := -⟨directed_system.Hid f, directed_system.Hcomp f⟩ +⟨directed_system.map_self f, directed_system.map_map f⟩ local attribute [instance] directed_system @@ -180,6 +180,7 @@ instance : add_comm_group (direct_limit G f) := module.direct_limit.add_comm_group G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) set_option class.instance_max_depth 50 + def of (i) : G i → direct_limit G f := module.direct_limit.of ℤ ι G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) i variables {G f} @@ -208,25 +209,25 @@ variables (g : Π i, G i → P) [Π i, is_add_group_hom (g i)] variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) variables (G f) -def rec : direct_limit G f → P := -module.direct_limit.rec ℤ ι G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) +def lift : direct_limit G f → P := +module.direct_limit.lift ℤ ι G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) (λ i, is_add_group_hom.to_linear_map $ g i) Hg variables {G f} -instance rec.is_add_group_hom : is_add_group_hom (rec G f P g Hg) := +instance lift.is_add_group_hom : is_add_group_hom (lift G f P g Hg) := linear_map.is_add_group_hom _ -@[simp] lemma rec_of (i x) : rec G f P g Hg (of G f i x) = g i x := -module.direct_limit.rec_of _ _ _ +@[simp] lemma lift_of (i x) : lift G f P g Hg (of G f i x) = g i x := +module.direct_limit.lift_of _ _ _ -@[simp] lemma rec_zero : rec G f P g Hg 0 = 0 := is_add_group_hom.zero _ -@[simp] lemma rec_add (x y) : rec G f P g Hg (x + y) = rec G f P g Hg x + rec G f P g Hg y := is_add_group_hom.add _ _ _ -@[simp] lemma rec_neg (x) : rec G f P g Hg (-x) = -rec G f P g Hg x := is_add_group_hom.neg _ _ -@[simp] lemma rec_sub (x y) : rec G f P g Hg (x - y) = rec G f P g Hg x - rec G f P g Hg y := is_add_group_hom.sub _ _ _ +@[simp] lemma lift_zero : lift G f P g Hg 0 = 0 := is_add_group_hom.zero _ +@[simp] lemma lift_add (x y) : lift G f P g Hg (x + y) = lift G f P g Hg x + lift G f P g Hg y := is_add_group_hom.add _ _ _ +@[simp] lemma lift_neg (x) : lift G f P g Hg (-x) = -lift G f P g Hg x := is_add_group_hom.neg _ _ +@[simp] lemma lift_sub (x y) : lift G f P g Hg (x - y) = lift G f P g Hg x - lift G f P g Hg y := is_add_group_hom.sub _ _ _ -lemma rec_unique (F : direct_limit G f → P) [is_add_group_hom F] (x) : - F x = rec G f P (λ i x, F $ of G f i x) (λ i j hij x, by rw of_f) x := -direct_limit.induction_on x $ λ i x, by rw rec_of +lemma lift_unique (F : direct_limit G f → P) [is_add_group_hom F] (x) : + F x = lift G f P (λ i x, F $ of G f i x) (λ i j hij x, by rw of_f) x := +direct_limit.induction_on x $ λ i x, by rw lift_of end direct_limit @@ -310,7 +311,7 @@ begin { rintros _ ⟨p, hps, rfl⟩ n ih, rw [restriction_mul, lift_mul, is_ring_hom.map_mul (f j k hjk), ih, restriction_mul, lift_mul, restriction_of, dif_pos hps, lift_pure, restriction_of, dif_pos (hst hps), lift_pure], - dsimp only, rw directed_system.Hcomp f, refl }, + dsimp only, rw directed_system.map_map f, refl }, { rintros x y ihx ihy, rw [restriction_add, lift_add, is_ring_hom.map_add (f j k hjk), ihx, ihy, restriction_add, lift_add] } end @@ -326,7 +327,7 @@ begin is_supported_sub (is_supported_pure.2 $ or.inl rfl) (is_supported_pure.2 $ or.inr $ or.inl rfl), _⟩, { rintros k (rfl | ⟨rfl | h⟩), refl, exact hij, cases h }, { rw [restriction_sub, lift_sub, restriction_of, dif_pos, restriction_of, dif_pos, lift_pure, lift_pure], - dsimp only, rw directed_system.Hcomp f, exact sub_self _, + dsimp only, rw directed_system.map_map f, exact sub_self _, { left, refl }, { right, left, refl }, } }, { refine ⟨i, {⟨i, 1⟩}, _, is_supported_sub (is_supported_pure.2 $ or.inl rfl) is_supported_one, _⟩, { rintros k (rfl | h), refl, cases h }, @@ -398,7 +399,7 @@ include Hg open free_comm_ring variables (G f) -def rec : direct_limit G f → P := +def lift : direct_limit G f → P := ideal.quotient.lift _ (free_comm_ring.lift $ λ x, g x.1 x.2) begin suffices : ideal.span _ ≤ ideal.comap (free_comm_ring.lift (λ (x : Σ (i : ι), G i), g (x.fst) (x.snd))) ⊥, @@ -412,23 +413,23 @@ end variables {G f} omit Hg -instance rec.is_ring_hom : is_ring_hom (rec G f P g Hg) := +instance lift.is_ring_hom : is_ring_hom (lift G f P g Hg) := ⟨free_comm_ring.lift_one _, λ x y, quotient.induction_on₂' x y $ λ p q, free_comm_ring.lift_mul _ _ _, λ x y, quotient.induction_on₂' x y $ λ p q, free_comm_ring.lift_add _ _ _⟩ -@[simp] lemma rec_of (i x) : rec G f P g Hg (of G f i x) = g i x := free_comm_ring.lift_of _ _ -@[simp] lemma rec_zero : rec G f P g Hg 0 = 0 := is_ring_hom.map_zero _ -@[simp] lemma rec_one : rec G f P g Hg 1 = 1 := is_ring_hom.map_one _ -@[simp] lemma rec_add (x y) : rec G f P g Hg (x + y) = rec G f P g Hg x + rec G f P g Hg y := is_ring_hom.map_add _ -@[simp] lemma rec_neg (x) : rec G f P g Hg (-x) = -rec G f P g Hg x := is_ring_hom.map_neg _ -@[simp] lemma rec_sub (x y) : rec G f P g Hg (x - y) = rec G f P g Hg x - rec G f P g Hg y := is_ring_hom.map_sub _ -@[simp] lemma rec_mul (x y) : rec G f P g Hg (x * y) = rec G f P g Hg x * rec G f P g Hg y := is_ring_hom.map_mul _ -@[simp] lemma rec_pow (x) (n : ℕ) : rec G f P g Hg (x ^ n) = rec G f P g Hg x ^ n := is_semiring_hom.map_pow _ _ _ - -theorem rec_unique (F : direct_limit G f → P) [is_ring_hom F] (x) : - F x = rec G f P (λ i x, F $ of G f i x) (λ i j hij x, by rw [of_f]) x := -direct_limit.induction_on x $ λ i x, by rw rec_of +@[simp] lemma lift_of (i x) : lift G f P g Hg (of G f i x) = g i x := free_comm_ring.lift_of _ _ +@[simp] lemma lift_zero : lift G f P g Hg 0 = 0 := is_ring_hom.map_zero _ +@[simp] lemma lift_one : lift G f P g Hg 1 = 1 := is_ring_hom.map_one _ +@[simp] lemma lift_add (x y) : lift G f P g Hg (x + y) = lift G f P g Hg x + lift G f P g Hg y := is_ring_hom.map_add _ +@[simp] lemma lift_neg (x) : lift G f P g Hg (-x) = -lift G f P g Hg x := is_ring_hom.map_neg _ +@[simp] lemma lift_sub (x y) : lift G f P g Hg (x - y) = lift G f P g Hg x - lift G f P g Hg y := is_ring_hom.map_sub _ +@[simp] lemma lift_mul (x y) : lift G f P g Hg (x * y) = lift G f P g Hg x * lift G f P g Hg y := is_ring_hom.map_mul _ +@[simp] lemma lift_pow (x) (n : ℕ) : lift G f P g Hg (x ^ n) = lift G f P g Hg x ^ n := is_semiring_hom.map_pow _ _ _ + +theorem lift_unique (F : direct_limit G f → P) [is_ring_hom F] (x) : + F x = lift G f P (λ i x, F $ of G f i x) (λ i j hij x, by rw [of_f]) x := +direct_limit.induction_on x $ λ i x, by rw lift_of end direct_limit From c01e2df44c58ee6817628b1fe12ccf67db6a999c Mon Sep 17 00:00:00 2001 From: kckennylau Date: Sun, 24 Feb 2019 14:02:24 +0000 Subject: [PATCH 18/22] docstrings --- src/algebra/direct_limit.lean | 40 +++++++++++++++++++++++++++++++++++ 1 file changed, 40 insertions(+) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 44d6ebe157f49..269bdf8f6596e 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -4,6 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Chris Hughes Direct limit of modules, abelian groups, rings, and fields. + +See Atiyah-Macdonald PP.32-33, Matsumura PP.269-270 + +Generalizes the notion of "union", or "gluing", of incomparable modules over the same ring, +or incomparable abelian groups, or rings, or fields. + +It is constructed as a quotient of the free module (for the module case) or quotient of +the free commutative ring (for the ring case) instead of a quotient of the disjoint union +so as to make the operations (addition etc.) "computable". -/ import linear_algebra.direct_sum_module @@ -20,6 +29,9 @@ variables {ι : Type v} [nonempty ι] variables [directed_order ι] [decidable_eq ι] variables (G : ι → Type w) [Π i, decidable_eq (G i)] +/-- A directed system is a functor from the category (directed poset) to another category. +This is used for abelian groups and rings and fields because their maps are not bundled. +See module.directed_system -/ class directed_system (f : Π i j, i ≤ j → G i → G j) : Prop := (map_self : ∀ i x h, f i i h x = x) (map_map : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) @@ -28,12 +40,14 @@ namespace module variables [Π i, add_comm_group (G i)] [Π i, module R (G i)] +/-- A directed system is a functor from the category (directed poset) to the category of R-modules. -/ class directed_system (f : Π i j, i ≤ j → G i →ₗ[R] G j) : Prop := (map_self : ∀ i x h, f i i h x = x) (map_map : ∀ i j k hij hjk x, f j k hjk (f i j hij x) = f i k (le_trans hij hjk) x) variables (f : Π i j, i ≤ j → G i →ₗ[R] G j) [directed_system G f] +/-- The direct limit of a directed system is the modules glued together along the maps. -/ def direct_limit : Type (max v w) := (span R $ { a | ∃ (i j) (H : i ≤ j) x, direct_sum.lof R ι G i x - direct_sum.lof R ι G j (f i j H x) = a }).quotient @@ -44,6 +58,7 @@ instance : add_comm_group (direct_limit G f) := quotient.add_comm_group _ instance : module R (direct_limit G f) := quotient.module _ variables (R ι) +/-- The canonical map from a component to the direct limit. -/ def of (i) : G i →ₗ[R] direct_limit G f := (mkq _).comp $ direct_sum.lof R ι G i variables {R ι G f} @@ -51,6 +66,8 @@ variables {R ι G f} @[simp] lemma of_f {i j hij x} : (of R ι G f j (f i j hij x)) = of R ι G f i x := eq.symm $ (submodule.quotient.eq _).2 $ subset_span ⟨i, j, hij, x, rfl⟩ +/-- Every element of the direct limit corresponds to some element in +some component of the directed system. -/ theorem exists_of (z : direct_limit G f) : ∃ i x, of R ι G f i x = z := nonempty.elim (by apply_instance) $ assume ind : ι, quotient.induction_on' z $ λ z, direct_sum.induction_on z @@ -69,6 +86,9 @@ variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) include Hg variables (R ι G f) +/-- The universal property of the direct limit: maps from the components to another module +that respect the directed system structure (i.e. make some diagram commute) give rise +to a unique map out of the direct limit. -/ def lift : direct_limit G f →ₗ[R] P := liftq _ (direct_sum.to_module R ι P g) (span_le.2 $ λ a ⟨i, j, hij, x, hx⟩, by rw [← hx, mem_coe, linear_map.sub_mem_ker_iff, @@ -142,6 +162,8 @@ span_induction ((quotient.mk_eq_zero _).1 H) ⟨i, λ k hk, hi k (dfinsupp.support_smul hk), by simp [linear_map.map_smul, hxi]⟩) +/-- A component that corresponds to zero in the direct limit is already zero in some +bigger module in the directed system. -/ theorem of.zero_exact {i x} (H : of R ι G f i x = 0) : ∃ j hij, f i j hij x = (0 : G j) := let ⟨j, hj, hxj⟩ := of.zero_exact_aux H in @@ -160,6 +182,7 @@ namespace add_comm_group variables [Π i, add_comm_group (G i)] +/-- The direct limit of a directed system is the abelian groups glued together along the maps. -/ def direct_limit (f : Π i j, i ≤ j → G i → G j) [Π i j hij, is_add_group_hom (f i j hij)] [directed_system G f] : Type* := @module.direct_limit ℤ _ ι _ _ _ G _ _ _ @@ -181,6 +204,7 @@ module.direct_limit.add_comm_group G (λ i j hij, is_add_group_hom.to_linear_map set_option class.instance_max_depth 50 +/-- The canonical map from a component to the direct limit. -/ def of (i) : G i → direct_limit G f := module.direct_limit.of ℤ ι G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) i variables {G f} @@ -201,6 +225,8 @@ protected theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit (ih : ∀ i x, C (of G f i x)) : C z := module.direct_limit.induction_on z ih +/-- A component that corresponds to zero in the direct limit is already zero in some +bigger module in the directed system. -/ theorem of.zero_exact (i x) (h : of G f i x = 0) : ∃ j hij, f i j hij x = 0 := module.direct_limit.of.zero_exact h @@ -209,6 +235,9 @@ variables (g : Π i, G i → P) [Π i, is_add_group_hom (g i)] variables (Hg : ∀ i j hij x, g j (f i j hij x) = g i x) variables (G f) +/-- The universal property of the direct limit: maps from the components to another abelian group +that respect the directed system structure (i.e. make some diagram commute) give rise +to a unique map out of the direct limit. -/ def lift : direct_limit G f → P := module.direct_limit.lift ℤ ι G (λ i j hij, is_add_group_hom.to_linear_map $ f i j hij) (λ i, is_add_group_hom.to_linear_map $ g i) Hg @@ -243,6 +272,7 @@ variables [directed_system G f] open free_comm_ring +/-- The direct limit of a directed system is the rings glued together along the maps. -/ def direct_limit : Type (max v w) := (ideal.span { a | (∃ i j H x, of (⟨j, f i j H x⟩ : Σ i, G i) - of ⟨i, x⟩ = a) ∨ (∃ i, of (⟨i, 1⟩ : Σ i, G i) - 1 = a) ∨ @@ -257,6 +287,7 @@ ideal.quotient.comm_ring _ instance : ring (direct_limit G f) := comm_ring.to_ring _ +/-- The canonical map from a component to the direct limit. -/ def of (i) (x : G i) : direct_limit G f := ideal.quotient.mk _ $ of ⟨i, x⟩ @@ -278,6 +309,8 @@ ideal.quotient.eq.2 $ subset_span $ or.inl ⟨i, j, hij, x, rfl⟩ @[simp] lemma of_mul (i x y) : of G f i (x * y) = of G f i x * of G f i y := is_ring_hom.map_mul _ @[simp] lemma of_pow (i x) (n : ℕ) : of G f i (x ^ n) = of G f i x ^ n := is_semiring_hom.map_pow _ _ _ +/-- Every element of the direct limit corresponds to some element in +some component of the directed system. -/ theorem exists_of (z : direct_limit G f) : ∃ i x, of G f i x = z := nonempty.elim (by apply_instance) $ assume ind : ι, quotient.induction_on' z $ λ x, free_abelian_group.induction_on x @@ -376,11 +409,15 @@ begin end end of_zero_exact_aux +/-- A component that corresponds to zero in the direct limit is already zero in some +bigger module in the directed system. -/ lemma of.zero_exact {i x} (hix : of G f i x = 0) : ∃ j, ∃ hij : i ≤ j, f i j hij x = 0 := let ⟨j, s, H, hxs, hx⟩ := of.zero_exact_aux hix in have hixs : (⟨i, x⟩ : Σ i, G i) ∈ s, from is_supported_pure.1 hxs, ⟨j, H ⟨i, x⟩ hixs, by rw [restriction_of, dif_pos hixs, lift_pure] at hx; exact hx⟩ +/-- If the maps in the directed system are injective, then the canonical maps +from the components to the direct limits are injective. -/ theorem of_inj (hf : ∀ i j hij, function.injective (f i j hij)) (i) : function.injective (of G f i) := begin @@ -399,6 +436,9 @@ include Hg open free_comm_ring variables (G f) +/-- The universal property of the direct limit: maps from the components to another ring +that respect the directed system structure (i.e. make some diagram commute) give rise +to a unique map out of the direct limit. -/ def lift : direct_limit G f → P := ideal.quotient.lift _ (free_comm_ring.lift $ λ x, g x.1 x.2) begin suffices : ideal.span _ ≤ From 2e99aa398da391816a7d6d4e10ca6dbae8e4a10a Mon Sep 17 00:00:00 2001 From: kckennylau Date: Sat, 25 May 2019 16:49:56 +0100 Subject: [PATCH 19/22] local --- src/algebra/direct_limit.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 269bdf8f6596e..534f40ef5216b 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -327,8 +327,8 @@ quotient.induction_on' z $ λ x, free_abelian_group.induction_on x (ih : ∀ i x, C (of G f i x)) : C z := let ⟨i, x, hx⟩ := exists_of z in hx ▸ ih i x -section of_zero_exact_aux -attribute [instance, priority 0] classical.dec +section of_zero_exact +local attribute [instance, priority 0] classical.dec variables (G f) lemma of.zero_exact_aux2 {x : free_comm_ring Σ i, G i} {s t} (hxs : is_supported x s) {j k} (hj : ∀ z : Σ i, G i, z ∈ s → z.1 ≤ j) (hk : ∀ z : Σ i, G i, z ∈ t → z.1 ≤ k) @@ -407,7 +407,6 @@ begin ← of.zero_exact_aux2 G f hyt hj this hjk (set.subset_union_right ↑s t), iht, is_ring_hom.map_zero (f j k hjk), mul_zero] } end -end of_zero_exact_aux /-- A component that corresponds to zero in the direct limit is already zero in some bigger module in the directed system. -/ @@ -415,6 +414,7 @@ lemma of.zero_exact {i x} (hix : of G f i x = 0) : ∃ j, ∃ hij : i ≤ j, f i let ⟨j, s, H, hxs, hx⟩ := of.zero_exact_aux hix in have hixs : (⟨i, x⟩ : Σ i, G i) ∈ s, from is_supported_pure.1 hxs, ⟨j, H ⟨i, x⟩ hixs, by rw [restriction_of, dif_pos hixs, lift_pure] at hx; exact hx⟩ +end of_zero_exact /-- If the maps in the directed system are injective, then the canonical maps from the components to the direct limits are injective. -/ From 3d8a480688ef2f96f014b6bf63994d5be023ed44 Mon Sep 17 00:00:00 2001 From: kckennylau Date: Thu, 30 May 2019 08:39:07 +0100 Subject: [PATCH 20/22] fix build --- src/algebra/direct_limit.lean | 42 +++++++++++++-------------- src/algebra/module.lean | 16 +++++----- src/ring_theory/free_comm_ring.lean | 45 +++++++++-------------------- src/ring_theory/free_ring.lean | 6 ++++ 4 files changed, 48 insertions(+), 61 deletions(-) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 534f40ef5216b..abda4a7ce8437 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -215,10 +215,10 @@ linear_map.is_add_group_hom _ @[simp] lemma of_f {i j} (hij) (x) : of G f j (f i j hij x) = of G f i x := module.direct_limit.of_f -@[simp] lemma of_zero (i) : of G f i 0 = 0 := is_add_group_hom.zero _ -@[simp] lemma of_add (i x y) : of G f i (x + y) = of G f i x + of G f i y := is_add_group_hom.add _ _ _ -@[simp] lemma of_neg (i x) : of G f i (-x) = -of G f i x := is_add_group_hom.neg _ _ -@[simp] lemma of_sub (i x y) : of G f i (x - y) = of G f i x - of G f i y := is_add_group_hom.sub _ _ _ +@[simp] lemma of_zero (i) : of G f i 0 = 0 := is_add_group_hom.map_zero _ +@[simp] lemma of_add (i x y) : of G f i (x + y) = of G f i x + of G f i y := is_add_group_hom.map_add _ _ _ +@[simp] lemma of_neg (i x) : of G f i (-x) = -of G f i x := is_add_group_hom.map_neg _ _ +@[simp] lemma of_sub (i x y) : of G f i (x - y) = of G f i x - of G f i y := is_add_group_hom.map_sub _ _ _ @[elab_as_eliminator] protected theorem induction_on {C : direct_limit G f → Prop} (z : direct_limit G f) @@ -249,10 +249,10 @@ linear_map.is_add_group_hom _ @[simp] lemma lift_of (i x) : lift G f P g Hg (of G f i x) = g i x := module.direct_limit.lift_of _ _ _ -@[simp] lemma lift_zero : lift G f P g Hg 0 = 0 := is_add_group_hom.zero _ -@[simp] lemma lift_add (x y) : lift G f P g Hg (x + y) = lift G f P g Hg x + lift G f P g Hg y := is_add_group_hom.add _ _ _ -@[simp] lemma lift_neg (x) : lift G f P g Hg (-x) = -lift G f P g Hg x := is_add_group_hom.neg _ _ -@[simp] lemma lift_sub (x y) : lift G f P g Hg (x - y) = lift G f P g Hg x - lift G f P g Hg y := is_add_group_hom.sub _ _ _ +@[simp] lemma lift_zero : lift G f P g Hg 0 = 0 := is_add_group_hom.map_zero _ +@[simp] lemma lift_add (x y) : lift G f P g Hg (x + y) = lift G f P g Hg x + lift G f P g Hg y := is_add_group_hom.map_add _ _ _ +@[simp] lemma lift_neg (x) : lift G f P g Hg (-x) = -lift G f P g Hg x := is_add_group_hom.map_neg _ _ +@[simp] lemma lift_sub (x y) : lift G f P g Hg (x - y) = lift G f P g Hg x - lift G f P g Hg y := is_add_group_hom.map_sub _ _ _ lemma lift_unique (F : direct_limit G f → P) [is_add_group_hom F] (x) : F x = lift G f P (λ i x, F $ of G f i x) (λ i j hij x, by rw of_f) x := @@ -343,7 +343,7 @@ begin restriction_neg, restriction_one, lift_neg, lift_one] }, { rintros _ ⟨p, hps, rfl⟩ n ih, rw [restriction_mul, lift_mul, is_ring_hom.map_mul (f j k hjk), ih, restriction_mul, lift_mul, - restriction_of, dif_pos hps, lift_pure, restriction_of, dif_pos (hst hps), lift_pure], + restriction_of, dif_pos hps, lift_of, restriction_of, dif_pos (hst hps), lift_of], dsimp only, rw directed_system.map_map f, refl }, { rintros x y ihx ihy, rw [restriction_add, lift_add, is_ring_hom.map_add (f j k hjk), ihx, ihy, restriction_add, lift_add] } @@ -357,29 +357,29 @@ begin refine span_induction (ideal.quotient.eq_zero_iff_mem.1 H) _ _ _ _, { rintros x (⟨i, j, hij, x, rfl⟩ | ⟨i, rfl⟩ | ⟨i, x, y, rfl⟩ | ⟨i, x, y, rfl⟩), { refine ⟨j, {⟨i, x⟩, ⟨j, f i j hij x⟩}, _, - is_supported_sub (is_supported_pure.2 $ or.inl rfl) (is_supported_pure.2 $ or.inr $ or.inl rfl), _⟩, + is_supported_sub (is_supported_of.2 $ or.inl rfl) (is_supported_of.2 $ or.inr $ or.inl rfl), _⟩, { rintros k (rfl | ⟨rfl | h⟩), refl, exact hij, cases h }, - { rw [restriction_sub, lift_sub, restriction_of, dif_pos, restriction_of, dif_pos, lift_pure, lift_pure], + { rw [restriction_sub, lift_sub, restriction_of, dif_pos, restriction_of, dif_pos, lift_of, lift_of], dsimp only, rw directed_system.map_map f, exact sub_self _, { left, refl }, { right, left, refl }, } }, - { refine ⟨i, {⟨i, 1⟩}, _, is_supported_sub (is_supported_pure.2 $ or.inl rfl) is_supported_one, _⟩, + { refine ⟨i, {⟨i, 1⟩}, _, is_supported_sub (is_supported_of.2 $ or.inl rfl) is_supported_one, _⟩, { rintros k (rfl | h), refl, cases h }, - { rw [restriction_sub, lift_sub, restriction_of, dif_pos, restriction_one, lift_pure, lift_one], + { rw [restriction_sub, lift_sub, restriction_of, dif_pos, restriction_one, lift_of, lift_one], dsimp only, rw [is_ring_hom.map_one (f i i _), sub_self], exact _inst_7 i i _, { left, refl } } }, { refine ⟨i, {⟨i, x+y⟩, ⟨i, x⟩, ⟨i, y⟩}, _, - is_supported_sub (is_supported_pure.2 $ or.inr $ or.inr $ or.inl rfl) - (is_supported_add (is_supported_pure.2 $ or.inr $ or.inl rfl) (is_supported_pure.2 $ or.inl rfl)), _⟩, + is_supported_sub (is_supported_of.2 $ or.inr $ or.inr $ or.inl rfl) + (is_supported_add (is_supported_of.2 $ or.inr $ or.inl rfl) (is_supported_of.2 $ or.inl rfl)), _⟩, { rintros k (rfl | ⟨rfl | ⟨rfl | hk⟩⟩), refl, refl, refl, cases hk }, { rw [restriction_sub, restriction_add, restriction_of, restriction_of, restriction_of, - dif_pos, dif_pos, dif_pos, lift_sub, lift_add, lift_pure, lift_pure, lift_pure], + dif_pos, dif_pos, dif_pos, lift_sub, lift_add, lift_of, lift_of, lift_of], dsimp only, rw is_ring_hom.map_add (f i i _), exact sub_self _, { right, right, left, refl }, { apply_instance }, { left, refl }, { right, left, refl } } }, { refine ⟨i, {⟨i, x*y⟩, ⟨i, x⟩, ⟨i, y⟩}, _, - is_supported_sub (is_supported_pure.2 $ or.inr $ or.inr $ or.inl rfl) - (is_supported_mul (is_supported_pure.2 $ or.inr $ or.inl rfl) (is_supported_pure.2 $ or.inl rfl)), _⟩, + is_supported_sub (is_supported_of.2 $ or.inr $ or.inr $ or.inl rfl) + (is_supported_mul (is_supported_of.2 $ or.inr $ or.inl rfl) (is_supported_of.2 $ or.inl rfl)), _⟩, { rintros k (rfl | ⟨rfl | ⟨rfl | hk⟩⟩), refl, refl, refl, cases hk }, { rw [restriction_sub, restriction_mul, restriction_of, restriction_of, restriction_of, - dif_pos, dif_pos, dif_pos, lift_sub, lift_mul, lift_pure, lift_pure, lift_pure], + dif_pos, dif_pos, dif_pos, lift_sub, lift_mul, lift_of, lift_of, lift_of], dsimp only, rw is_ring_hom.map_mul (f i i _), exact sub_self _, { right, right, left, refl }, { apply_instance }, { left, refl }, { right, left, refl } } } }, { refine nonempty.elim (by apply_instance) (assume ind : ι, _), @@ -412,8 +412,8 @@ end bigger module in the directed system. -/ lemma of.zero_exact {i x} (hix : of G f i x = 0) : ∃ j, ∃ hij : i ≤ j, f i j hij x = 0 := let ⟨j, s, H, hxs, hx⟩ := of.zero_exact_aux hix in -have hixs : (⟨i, x⟩ : Σ i, G i) ∈ s, from is_supported_pure.1 hxs, -⟨j, H ⟨i, x⟩ hixs, by rw [restriction_of, dif_pos hixs, lift_pure] at hx; exact hx⟩ +have hixs : (⟨i, x⟩ : Σ i, G i) ∈ s, from is_supported_of.1 hxs, +⟨j, H ⟨i, x⟩ hixs, by rw [restriction_of, dif_pos hixs, lift_of] at hx; exact hx⟩ end of_zero_exact /-- If the maps in the directed system are injective, then the canonical maps diff --git a/src/algebra/module.lean b/src/algebra/module.lean index 7a7ed77886ea1..772c6157603bb 100644 --- a/src/algebra/module.lean +++ b/src/algebra/module.lean @@ -169,14 +169,6 @@ def id : β →ₗ[α] β := ⟨id, by simp, by simp⟩ end linear_map -def is_add_group_hom.to_linear_map [add_comm_group α] [add_comm_group β] - (f : α → β) [is_add_group_hom f] : α →ₗ[ℤ] β := -{ to_fun := f, - add := is_add_group_hom.add f, - smul := λ i x, int.induction_on i (by rw [zero_smul, zero_smul, is_add_group_hom.zero f]) - (λ i ih, by rw [add_smul, add_smul, is_add_group_hom.add f, ih, one_smul, one_smul]) - (λ i ih, by rw [sub_smul, sub_smul, is_add_group_hom.sub f, ih, one_smul, one_smul]) } - namespace is_linear_map variables [ring α] [add_comm_group β] [add_comm_group γ] variables [module α β] [module α γ] @@ -398,3 +390,11 @@ instance : module ℤ M := smul_zero := gsmul_zero } end add_comm_group + +def is_add_group_hom.to_linear_map [add_comm_group α] [add_comm_group β] + (f : α → β) [is_add_group_hom f] : α →ₗ[ℤ] β := +{ to_fun := f, + add := is_add_group_hom.map_add f, + smul := λ i x, int.induction_on i (by rw [zero_smul, zero_smul, is_add_group_hom.map_zero f]) + (λ i ih, by rw [add_smul, add_smul, is_add_group_hom.map_add f, ih, one_smul, one_smul]) + (λ i ih, by rw [sub_smul, sub_smul, is_add_group_hom.map_sub f, ih, one_smul, one_smul]) } diff --git a/src/ring_theory/free_comm_ring.lean b/src/ring_theory/free_comm_ring.lean index 55b89b4cb511b..9ba8c7de5c85a 100644 --- a/src/ring_theory/free_comm_ring.lean +++ b/src/ring_theory/free_comm_ring.lean @@ -1,10 +1,11 @@ /- Copyright (c) 2019 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kenny Lau +Authors: Kenny Lau, Johan Commelin -/ import group_theory.free_abelian_group data.equiv.algebra data.equiv.functor data.polynomial +import data.real.irrational import ring_theory.ideal_operations ring_theory.free_ring universes u v @@ -34,7 +35,6 @@ free_abelian_group.induction_on z (λ m, multiset.induction_on m h1 $ λ a m ih, hm _ _ (hb a) ih) (λ m ih, hn _ ih) ha - section lift variables {β : Type v} [comm_ring β] (f : α → β) @@ -97,23 +97,6 @@ variables {β : Type v} (f : α → β) def map : free_comm_ring α → free_comm_ring β := lift $ of ∘ f -instance : monad free_comm_ring := -{ map := λ _ _, map, - pure := λ _, of, - bind := λ _ _ x f, lift f x } - -@[simp] lemma lift_pure [comm_ring β] (x : α) : lift f (pure x) = f x := -lift_of _ _ - -@[elab_as_eliminator] protected lemma induction_on' - {C : free_comm_ring α → Prop} (z : free_comm_ring α) - (hn1 : C (-1)) (hb : ∀ b, C (pure b)) - (ha : ∀ x y, C x → C y → C (x + y)) - (hm : ∀ x y, C x → C y → C (x * y)) : C z := -free_comm_ring.induction_on z hn1 hb ha hm - -section map -@[simp] lemma map_pure (z) : f <$> (pure z : free_comm_ring α) = pure (f z) := lift_of (of ∘ f) @[simp] lemma map_zero : map f 0 = 0 := rfl @[simp] lemma map_one : map f 1 = 1 := rfl @[simp] lemma map_of (x : α) : map f (of x) = of (f x) := lift_of _ _ @@ -122,8 +105,6 @@ section map @[simp] lemma map_sub (x y) : map f (x - y) = map f x - map f y := lift_sub _ _ _ @[simp] lemma map_mul (x y) : map f (x * y) = map f x * map f y := lift_mul _ _ _ @[simp] lemma map_pow (x) (n : ℕ) : map f (x ^ n) = (map f x) ^ n := lift_pow _ _ _ -end map - def is_supported (x : free_comm_ring α) (s : set α) : Prop := x ∈ ring.closure (of '' s) @@ -160,12 +141,11 @@ int.induction_on i is_supported_zero (λ i hi, by rw [int.cast_sub, int.cast_one]; exact is_supported_sub hi is_supported_one) def restriction (s : set α) [decidable_pred s] (x : free_comm_ring α) : free_comm_ring s := -lift (λ p, if H : p ∈ s then pure ⟨p, H⟩ else 0) x +lift (λ p, if H : p ∈ s then of ⟨p, H⟩ else 0) x section restriction variables (s : set α) [decidable_pred s] (x y : free_comm_ring α) -@[simp] lemma restriction_of (p) : restriction s (of p) = if H : p ∈ s then pure ⟨p, H⟩ else 0 := lift_of _ _ -@[simp] lemma restriction_pure (p) : restriction s (pure p) = if H : p ∈ s then pure ⟨p, H⟩ else 0 := lift_of _ _ +@[simp] lemma restriction_of (p) : restriction s (of p) = if H : p ∈ s then of ⟨p, H⟩ else 0 := lift_of _ _ @[simp] lemma restriction_zero : restriction s 0 = 0 := lift_zero _ @[simp] lemma restriction_one : restriction s 1 = 1 := lift_one _ @[simp] lemma restriction_add : restriction s (x + y) = restriction s x + restriction s y := lift_add _ _ _ @@ -174,34 +154,34 @@ variables (s : set α) [decidable_pred s] (x y : free_comm_ring α) @[simp] lemma restriction_mul : restriction s (x * y) = restriction s x * restriction s y := lift_mul _ _ _ end restriction -theorem is_supported_pure {p} {s : set α} : is_supported (pure p) s ↔ p ∈ s := -suffices is_supported (pure p) s → p ∈ s, from ⟨this, λ hps, ring.subset_closure ⟨p, hps, rfl⟩⟩, -assume hps : is_supported (pure p) s, begin +theorem is_supported_of {p} {s : set α} : is_supported (of p) s ↔ p ∈ s := +suffices is_supported (of p) s → p ∈ s, from ⟨this, λ hps, ring.subset_closure ⟨p, hps, rfl⟩⟩, +assume hps : is_supported (of p) s, begin classical, have : ∀ x, is_supported x s → ∃ q : ℚ, lift (λ q, if q ∈ s then 0 else real.sqrt 2) x = ↑q, { intros x hx, refine ring.in_closure.rec_on hx _ _ _ _, { use 1, rw [lift_one, rat.cast_one] }, { use -1, rw [lift_neg, lift_one, rat.cast_neg, rat.cast_one], }, { rintros _ ⟨z, hzs, rfl⟩ _ _, use 0, rw [lift_mul, lift_of, if_pos hzs, zero_mul, rat.cast_zero] }, - { rintros x y ⟨q, hq⟩ ⟨r, hr⟩, use q+r, rw [lift_add, hq, hr, rat.cast_add] } }, + { rintros x y ⟨q, hq⟩ ⟨r, hr⟩, refine ⟨q+r, _⟩, rw [lift_add, hq, hr, rat.cast_add] } }, specialize this (of p) hps, rw [lift_of] at this, split_ifs at this, { exact h }, exfalso, exact irr_sqrt_two this end -theorem subtype_val_map_restriction {x} (s : set α) [decidable_pred s] (hxs : is_supported x s) : - subtype.val <$> restriction s x = x := +theorem map_subtype_val_restriction {x} (s : set α) [decidable_pred s] (hxs : is_supported x s) : + map subtype.val (restriction s x) = x := begin refine ring.in_closure.rec_on hxs _ _ _ _, { rw restriction_one, refl }, { rw [restriction_neg, map_neg, restriction_one], refl }, - { rintros _ ⟨p, hps, rfl⟩ n ih, rw [restriction_mul, restriction_of, dif_pos hps, map_mul, map_pure, ih], refl }, + { rintros _ ⟨p, hps, rfl⟩ n ih, rw [restriction_mul, restriction_of, dif_pos hps, map_mul, map_of, ih] }, { intros x y ihx ihy, rw [restriction_add, map_add, ihx, ihy] } end theorem exists_finite_support (x : free_comm_ring α) : ∃ s : set α, set.finite s ∧ is_supported x s := free_comm_ring.induction_on x ⟨∅, set.finite_empty, is_supported_neg is_supported_one⟩ - (λ p, ⟨{p}, set.finite_singleton p, is_supported_pure.2 $ finset.mem_singleton_self _⟩) + (λ p, ⟨{p}, set.finite_singleton p, is_supported_of.2 $ finset.mem_singleton_self _⟩) (λ x y ⟨s, hfs, hxs⟩ ⟨t, hft, hxt⟩, ⟨s ∪ t, set.finite_union hfs hft, is_supported_add (is_supported_upwards hxs $ set.subset_union_left s t) (is_supported_upwards hxt $ set.subset_union_right s t)⟩) @@ -214,6 +194,7 @@ let ⟨s, hfs, hxs⟩ := exists_finite_support x in ⟨hfs.to_finset, by rwa fin end free_comm_ring + namespace free_ring open function variable (α) diff --git a/src/ring_theory/free_ring.lean b/src/ring_theory/free_ring.lean index 5940d40915c51..65704a5f1d7a4 100644 --- a/src/ring_theory/free_ring.lean +++ b/src/ring_theory/free_ring.lean @@ -1,3 +1,9 @@ +/- +Copyright (c) 2019 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau, Johan Commelin +-/ + import group_theory.free_abelian_group data.equiv.algebra data.polynomial universes u v From 85937c6bf62281e59d5a87b85a3fad7eba977c1c Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 31 May 2019 16:25:45 +0200 Subject: [PATCH 21/22] Replace real with polynomial int in proof --- src/data/polynomial.lean | 12 ++++++++++ src/ring_theory/free_comm_ring.lean | 35 +++++++++++++++++------------ 2 files changed, 33 insertions(+), 14 deletions(-) diff --git a/src/data/polynomial.lean b/src/data/polynomial.lean index 8d37bfd8bf0ff..a47d0c3947c5c 100644 --- a/src/data/polynomial.lean +++ b/src/data/polynomial.lean @@ -116,6 +116,9 @@ instance C.is_semiring_hom : is_semiring_hom (C : α → polynomial α) := @[simp] lemma C_pow : C (a ^ n) = C a ^ n := is_semiring_hom.map_pow _ _ _ +lemma nat_cast_eq_C (n : ℕ) : (n : polynomial α) = C n := +by refine (nat.eq_cast (λ n, C (n : α)) _ _ _ n).symm; simp + section coeff lemma apply_eq_coeff : p n = coeff p n := rfl @@ -518,6 +521,9 @@ end @[simp] lemma nat_degree_one : nat_degree (1 : polynomial α) = 0 := nat_degree_C 1 +@[simp] lemma nat_degree_nat_cast (n : ℕ) : nat_degree (n : polynomial α) = 0 := +by simp [nat_cast_eq_C] + @[simp] lemma degree_monomial (n : ℕ) (ha : a ≠ 0) : degree (C a * X ^ n) = n := by rw [← single_eq_C_mul_X, degree, support_single_ne_zero ha]; refl @@ -1141,6 +1147,9 @@ variable {α} instance C.is_ring_hom : is_ring_hom (@C α _ _) := by apply is_ring_hom.of_semiring +lemma int_cast_eq_C (n : ℤ) : (n : polynomial α) = C n := +congr_fun (int.eq_cast' _).symm n + @[simp] lemma C_neg : C (-a) = -C a := is_ring_hom.map_neg C @[simp] lemma C_sub : C (a - b) = C a - C b := is_ring_hom.map_sub C @@ -1164,6 +1173,9 @@ eval₂.is_ring_hom (C ∘ f) @[simp] lemma degree_neg (p : polynomial α) : degree (-p) = degree p := by unfold degree; rw support_neg +@[simp] lemma nat_degree_int_cast (n : ℤ) : nat_degree (n : polynomial α) = 0 := +by simp [int_cast_eq_C] + @[simp] lemma coeff_neg (p : polynomial α) (n : ℕ) : coeff (-p) n = -coeff p n := rfl @[simp] lemma coeff_sub (p q : polynomial α) (n : ℕ) : coeff (p - q) n = coeff p n - coeff q n := rfl diff --git a/src/ring_theory/free_comm_ring.lean b/src/ring_theory/free_comm_ring.lean index 9ba8c7de5c85a..231bb3e87aaa2 100644 --- a/src/ring_theory/free_comm_ring.lean +++ b/src/ring_theory/free_comm_ring.lean @@ -5,7 +5,6 @@ Authors: Kenny Lau, Johan Commelin -/ import group_theory.free_abelian_group data.equiv.algebra data.equiv.functor data.polynomial -import data.real.irrational import ring_theory.ideal_operations ring_theory.free_ring universes u v @@ -109,30 +108,33 @@ lift $ of ∘ f def is_supported (x : free_comm_ring α) (s : set α) : Prop := x ∈ ring.closure (of '' s) -theorem is_supported_upwards {x : free_comm_ring α} {s t} (hs : is_supported x s) (hst : s ⊆ t) : +section is_supported +variables {x y : free_comm_ring α} {s t : set α} + +theorem is_supported_upwards (hs : is_supported x s) (hst : s ⊆ t) : is_supported x t := ring.closure_mono (set.mono_image hst) hs -theorem is_supported_add {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : +theorem is_supported_add (hxs : is_supported x s) (hys : is_supported y s) : is_supported (x + y) s := is_add_submonoid.add_mem hxs hys -theorem is_supported_neg {x : free_comm_ring α} {s} (hxs : is_supported x s) : +theorem is_supported_neg (hxs : is_supported x s) : is_supported (-x) s := is_add_subgroup.neg_mem hxs -theorem is_supported_sub {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : +theorem is_supported_sub (hxs : is_supported x s) (hys : is_supported y s) : is_supported (x - y) s := is_add_subgroup.sub_mem _ _ _ hxs hys -theorem is_supported_mul {x y : free_comm_ring α} {s} (hxs : is_supported x s) (hys : is_supported y s) : +theorem is_supported_mul (hxs : is_supported x s) (hys : is_supported y s) : is_supported (x * y) s := is_submonoid.mul_mem hxs hys -theorem is_supported_zero {s : set α} : is_supported 0 s := +theorem is_supported_zero : is_supported 0 s := is_add_submonoid.zero_mem _ -theorem is_supported_one {s : set α} : is_supported 1 s := +theorem is_supported_one : is_supported 1 s := is_submonoid.one_mem _ theorem is_supported_int {i : ℤ} {s : set α} : is_supported ↑i s := @@ -140,6 +142,8 @@ int.induction_on i is_supported_zero (λ i hi, by rw [int.cast_add, int.cast_one]; exact is_supported_add hi is_supported_one) (λ i hi, by rw [int.cast_sub, int.cast_one]; exact is_supported_sub hi is_supported_one) +end is_supported + def restriction (s : set α) [decidable_pred s] (x : free_comm_ring α) : free_comm_ring s := lift (λ p, if H : p ∈ s then of ⟨p, H⟩ else 0) x @@ -158,14 +162,17 @@ theorem is_supported_of {p} {s : set α} : is_supported (of p) s ↔ p ∈ s := suffices is_supported (of p) s → p ∈ s, from ⟨this, λ hps, ring.subset_closure ⟨p, hps, rfl⟩⟩, assume hps : is_supported (of p) s, begin classical, - have : ∀ x, is_supported x s → ∃ q : ℚ, lift (λ q, if q ∈ s then 0 else real.sqrt 2) x = ↑q, + have : ∀ x, is_supported x s → + ∃ (n : ℤ), lift (λ a, if a ∈ s then (0 : polynomial ℤ) else polynomial.X) x = n, { intros x hx, refine ring.in_closure.rec_on hx _ _ _ _, - { use 1, rw [lift_one, rat.cast_one] }, - { use -1, rw [lift_neg, lift_one, rat.cast_neg, rat.cast_one], }, - { rintros _ ⟨z, hzs, rfl⟩ _ _, use 0, rw [lift_mul, lift_of, if_pos hzs, zero_mul, rat.cast_zero] }, - { rintros x y ⟨q, hq⟩ ⟨r, hr⟩, refine ⟨q+r, _⟩, rw [lift_add, hq, hr, rat.cast_add] } }, + { use 1, rw [lift_one], norm_cast }, + { use -1, rw [lift_neg, lift_one], norm_cast }, + { rintros _ ⟨z, hzs, rfl⟩ _ _, use 0, rw [lift_mul, lift_of, if_pos hzs, zero_mul], norm_cast }, + { rintros x y ⟨q, hq⟩ ⟨r, hr⟩, refine ⟨q+r, _⟩, rw [lift_add, hq, hr], norm_cast } }, specialize this (of p) hps, rw [lift_of] at this, split_ifs at this, { exact h }, - exfalso, exact irr_sqrt_two this + exfalso, apply int.zero_ne_one, + rcases this with ⟨w, H⟩, rw polynomial.int_cast_eq_C at H, + exact congr_arg (λ (f : polynomial ℤ), f.coeff 1) H.symm end theorem map_subtype_val_restriction {x} (s : set α) [decidable_pred s] (hxs : is_supported x s) : From 322c273a474ee798671055ca3a7b39966ee87d46 Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Sun, 16 Jun 2019 18:37:47 +0100 Subject: [PATCH 22/22] Update basic.lean --- src/order/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/order/basic.lean b/src/order/basic.lean index e9ed9e3b5b180..a9085576aefe0 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -538,5 +538,5 @@ theorem directed_mono {s : α → α → Prop} {ι} (f : ι → α) (H : ∀ a b, r a b → s a b) (h : directed r f) : directed s f := λ a b, let ⟨c, h₁, h₂⟩ := h a b in ⟨c, H _ _ h₁, H _ _ h₂⟩ -class directed_order (α : Type u) extends partial_order α := +class directed_order (α : Type u) extends preorder α := (directed : ∀ i j : α, ∃ k, i ≤ k ∧ j ≤ k)