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".
- map_self : ∀ (i : ι) (x : G i) (h : i ≤ i), f i i h x = x
- map_map : ∀ (i j k : ι) (hij : i ≤ j) (hjk : j ≤ k) (x : G i), f j k hjk (f i j hij x) = f i k _ x
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
- map_self : ∀ (i : ι) (x : G i) (h : i ≤ i), ⇑(f i i h) x = x
- map_map : ∀ (i j k : ι) (hij : i ≤ j) (hjk : j ≤ k) (x : G i), ⇑(f j k hjk) (⇑(f i j hij) x) = ⇑(f i k _) x
A directed system is a functor from the category (directed poset) to the category of R-modules.
The direct limit of a directed system is the modules glued together along the maps.
Equations
- module.direct_limit G f = (submodule.span R {a : direct_sum ι (λ (i : ι), G i) | ∃ (i j : ι) (H : i ≤ j) (x : G i), ⇑(direct_sum.lof R ι G i) x - ⇑(direct_sum.lof R ι G j) (⇑(f i j H) x) = a}).quotient
Equations
- module.direct_limit.add_comm_group G f = submodule.quotient.add_comm_group (submodule.span R {a : direct_sum ι (λ (i : ι), G i) | ∃ (i j : ι) (H : i ≤ j) (x : G i), ⇑(direct_sum.lof R ι G i) x - ⇑(direct_sum.lof R ι G j) (⇑(f i j H) x) = a})
Equations
- module.direct_limit.semimodule G f = submodule.quotient.semimodule (submodule.span R {a : direct_sum ι (λ (i : ι), G i) | ∃ (i j : ι) (H : i ≤ j) (x : G i), ⇑(direct_sum.lof R ι G i) x - ⇑(direct_sum.lof R ι G j) (⇑(f i j H) x) = a})
The canonical map from a component to the direct limit.
Equations
- module.direct_limit.of R ι G f i = (submodule.span R {a : direct_sum ι (λ (i : ι), G i) | ∃ (i j : ι) (H : i ≤ j) (x : G i), ⇑(direct_sum.lof R ι G i) x - ⇑(direct_sum.lof R ι G j) (⇑(f i j H) x) = a}).mkq.comp (direct_sum.lof R ι G i)
Every element of the direct limit corresponds to some element in some component of the directed system.
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.
Equations
- module.direct_limit.lift R ι G f g Hg = (submodule.span R {a : direct_sum ι (λ (i : ι), G i) | ∃ (i j : ι) (H : i ≤ j) (x : G i), ⇑(direct_sum.lof R ι G i) x - ⇑(direct_sum.lof R ι G j) (⇑(f i j H) x) = a}).liftq (direct_sum.to_module R ι P g) _
A component that corresponds to zero in the direct limit is already zero in some bigger module in the directed system.
The direct limit of a directed system is the abelian groups glued together along the maps.
Equations
- add_comm_group.direct_limit G f = module.direct_limit G (λ (i j : ι) (hij : i ≤ j), (add_monoid_hom.of (f i j hij)).to_int_linear_map)
Equations
- add_comm_group.direct_limit.add_comm_group G f = module.direct_limit.add_comm_group G (λ (i j : ι) (hij : i ≤ j), (add_monoid_hom.of (f i j hij)).to_int_linear_map)
The canonical map from a component to the direct limit.
Equations
- add_comm_group.direct_limit.of G f i = ⇑(module.direct_limit.of ℤ ι G (λ (i j : ι) (hij : i ≤ j), (add_monoid_hom.of (f i j hij)).to_int_linear_map) i)
Equations
- _ = _
A component that corresponds to zero in the direct limit is already zero in some bigger module in the directed system.
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.
Equations
- add_comm_group.direct_limit.lift G f P g Hg = ⇑(module.direct_limit.lift ℤ ι G (λ (i j : ι) (hij : i ≤ j), (add_monoid_hom.of (f i j hij)).to_int_linear_map) (λ (i : ι), (add_monoid_hom.of (g i)).to_int_linear_map) Hg)
Equations
- _ = _
The direct limit of a directed system is the rings glued together along the maps.
Equations
- ring.direct_limit G f = (ideal.span {a : free_comm_ring (Σ (i : ι), G i) | (∃ (i j : ι) (H : i ≤ j) (x : G i), free_comm_ring.of ⟨j, f i j H x⟩ - free_comm_ring.of ⟨i, x⟩ = a) ∨ (∃ (i : ι), free_comm_ring.of ⟨i, 1⟩ - 1 = a) ∨ (∃ (i : ι) (x y : G i), free_comm_ring.of ⟨i, x + y⟩ - (free_comm_ring.of ⟨i, x⟩ + free_comm_ring.of ⟨i, y⟩) = a) ∨ ∃ (i : ι) (x y : G i), free_comm_ring.of ⟨i, x * y⟩ - free_comm_ring.of ⟨i, x⟩ * free_comm_ring.of ⟨i, y⟩ = a}).quotient
Equations
- ring.direct_limit.comm_ring G f = ideal.quotient.comm_ring (ideal.span {a : free_comm_ring (Σ (i : ι), G i) | (∃ (i j : ι) (H : i ≤ j) (x : G i), free_comm_ring.of ⟨j, f i j H x⟩ - free_comm_ring.of ⟨i, x⟩ = a) ∨ (∃ (i : ι), free_comm_ring.of ⟨i, 1⟩ - 1 = a) ∨ (∃ (i : ι) (x y : G i), free_comm_ring.of ⟨i, x + y⟩ - (free_comm_ring.of ⟨i, x⟩ + free_comm_ring.of ⟨i, y⟩) = a) ∨ ∃ (i : ι) (x y : G i), free_comm_ring.of ⟨i, x * y⟩ - free_comm_ring.of ⟨i, x⟩ * free_comm_ring.of ⟨i, y⟩ = a})
Equations
The canonical map from a component to the direct limit.
Equations
- ring.direct_limit.of G f i x = ⇑(ideal.quotient.mk (ideal.span {a : free_comm_ring (Σ (i : ι), G i) | (∃ (i j : ι) (H : i ≤ j) (x : G i), free_comm_ring.of ⟨j, f i j H x⟩ - free_comm_ring.of ⟨i, x⟩ = a) ∨ (∃ (i : ι), free_comm_ring.of ⟨i, 1⟩ - 1 = a) ∨ (∃ (i : ι) (x y : G i), free_comm_ring.of ⟨i, x + y⟩ - (free_comm_ring.of ⟨i, x⟩ + free_comm_ring.of ⟨i, y⟩) = a) ∨ ∃ (i : ι) (x y : G i), free_comm_ring.of ⟨i, x * y⟩ - free_comm_ring.of ⟨i, x⟩ * free_comm_ring.of ⟨i, y⟩ = a})) (free_comm_ring.of ⟨i, x⟩)
Equations
- _ = _
Every element of the direct limit corresponds to some element in some component of the directed system.
A component that corresponds to zero in the direct limit is already zero in some bigger module in the directed system.
If the maps in the directed system are injective, then the canonical maps from the components to the direct limits are injective.
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.
We don't use this function as the canonical form because Lean 3 fails to automatically coerce
it to a function; use lift
instead.
Equations
- ring.direct_limit.lift_hom G f P g Hg = ideal.quotient.lift (ideal.span {a : free_comm_ring (Σ (i : ι), G i) | (∃ (i j : ι) (H : i ≤ j) (x : G i), free_comm_ring.of ⟨j, f i j H x⟩ - free_comm_ring.of ⟨i, x⟩ = a) ∨ (∃ (i : ι), free_comm_ring.of ⟨i, 1⟩ - 1 = a) ∨ (∃ (i : ι) (x y : G i), free_comm_ring.of ⟨i, x + y⟩ - (free_comm_ring.of ⟨i, x⟩ + free_comm_ring.of ⟨i, y⟩) = a) ∨ ∃ (i : ι) (x y : G i), free_comm_ring.of ⟨i, x * y⟩ - free_comm_ring.of ⟨i, x⟩ * free_comm_ring.of ⟨i, y⟩ = a}) (free_comm_ring.lift_hom (λ (x : Σ (i : ι), G i), g x.fst x.snd)) _
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.
Equations
- ring.direct_limit.lift G f P g Hg = ⇑(ring.direct_limit.lift_hom G f P g Hg)
Equations
- _ = _
Equations
- _ = _
Equations
- field.direct_limit.inv G f p = dite (p = 0) (λ (H : p = 0), 0) (λ (H : ¬p = 0), classical.some _)
Equations
- field.direct_limit.field G f = {add := comm_ring.add (ring.direct_limit.comm_ring G f), add_assoc := _, zero := comm_ring.zero (ring.direct_limit.comm_ring G f), zero_add := _, add_zero := _, neg := comm_ring.neg (ring.direct_limit.comm_ring G f), add_left_neg := _, add_comm := _, mul := comm_ring.mul (ring.direct_limit.comm_ring G f), mul_assoc := _, one := comm_ring.one (ring.direct_limit.comm_ring G f), one_mul := _, mul_one := _, left_distrib := _, right_distrib := _, mul_comm := _, inv := field.direct_limit.inv G f _inst_8, exists_pair_ne := _, mul_inv_cancel := _, inv_zero := _}