|
| 1 | +/- |
| 2 | +Copyright (c) 2020 Johan Commelin. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Johan Commelin, Robert Y. Lewis |
| 5 | +-/ |
| 6 | + |
| 7 | +import data.mv_polynomial.counit |
| 8 | +import data.mv_polynomial.invertible |
| 9 | +import ring_theory.witt_vector.defs |
| 10 | + |
| 11 | +/-! |
| 12 | +# Witt vectors |
| 13 | +
|
| 14 | +This file verifies that the ring operations on `witt_vector p R` |
| 15 | +satisfy the axioms of a commutative ring. |
| 16 | +
|
| 17 | +## Main definitions |
| 18 | +
|
| 19 | +* `witt_vector.map`: lifts a ring homomorphism `R →+* S` to a ring homomorphism `𝕎 R →+* 𝕎 S`. |
| 20 | +* `witt_vector.ghost_component n x`: evaluates the `n`th Witt polynomial |
| 21 | + on the first `n` coefficients of `x`, producing a value in `R`. |
| 22 | + This is a ring homomorphism. |
| 23 | +* `witt_vector.ghost_map`: a ring homomorphism `𝕎 R →+* (ℕ → R)`, obtained by packaging |
| 24 | + all the ghost components together. |
| 25 | + If `p` is invertible in `R`, then the ghost map is an equivalence, |
| 26 | + which we use to define the ring operations on `𝕎 R`. |
| 27 | +* `witt_vector.comm_ring`: the ring structure induced by the ghost components. |
| 28 | +
|
| 29 | +## Notation |
| 30 | +
|
| 31 | +We use notation `𝕎 R`, entered `\bbW`, for the Witt vectors over `R`. |
| 32 | +
|
| 33 | +## Implementation details |
| 34 | +
|
| 35 | +As we prove that the ghost components respect the ring operations, we face a number of repetitive |
| 36 | +proofs. To avoid duplicating code we factor these proofs into a custom tactic, only slightly more |
| 37 | +powerful than a tactic macro. This tactic is not particularly useful outside of its applications |
| 38 | +in this file. |
| 39 | +-/ |
| 40 | + |
| 41 | +noncomputable theory |
| 42 | + |
| 43 | +open mv_polynomial function |
| 44 | + |
| 45 | +open_locale big_operators |
| 46 | + |
| 47 | +local attribute [semireducible] witt_vector |
| 48 | + |
| 49 | +variables {p : ℕ} {R S T : Type*} [hp : fact p.prime] [comm_ring R] [comm_ring S] [comm_ring T] |
| 50 | +variables {α : Type*} {β : Type*} |
| 51 | + |
| 52 | +local notation `𝕎` := witt_vector p -- type as `\bbW` |
| 53 | +open_locale witt |
| 54 | + |
| 55 | +namespace witt_vector |
| 56 | + |
| 57 | +/-- `f : α → β` induces a map from `𝕎 α` to `𝕎 β` by applying `f` componentwise. |
| 58 | +If `f` is a ring homomorphism, then so is `f`, see `witt_vector.map f`. -/ |
| 59 | +def map_fun (f : α → β) : 𝕎 α → 𝕎 β := λ x, f ∘ x |
| 60 | + |
| 61 | +namespace map_fun |
| 62 | + |
| 63 | +lemma injective (f : α → β) (hf : injective f) : injective (map_fun f : 𝕎 α → 𝕎 β) := |
| 64 | +λ x y h, funext $ λ n, hf $ by exact congr_fun h n |
| 65 | + |
| 66 | +lemma surjective (f : α → β) (hf : surjective f) : surjective (map_fun f : 𝕎 α → 𝕎 β) := |
| 67 | +λ x, ⟨λ n, classical.some $ hf $ x n, |
| 68 | +by { funext n, dsimp [map_fun], rw classical.some_spec (hf (x n)) }⟩ |
| 69 | + |
| 70 | +variables (f : R →+* S) (x y : 𝕎 R) |
| 71 | + |
| 72 | +/-- Auxiliary tactic for showing that `map_fun` respects the ring operations. -/ |
| 73 | +meta def map_fun_tac : tactic unit := |
| 74 | +`[funext n, |
| 75 | + show f (aeval _ _) = aeval _ _, |
| 76 | + rw map_aeval, |
| 77 | + apply eval₂_hom_congr (ring_hom.ext_int _ _) _ rfl, |
| 78 | + ext ⟨i, k⟩, |
| 79 | + fin_cases i; refl] |
| 80 | + |
| 81 | +include hp |
| 82 | + |
| 83 | +/- We do not tag these lemmas as `@[simp]` because they will be bundled in `map` later on. -/ |
| 84 | + |
| 85 | +lemma zero : map_fun f (0 : 𝕎 R) = 0 := by map_fun_tac |
| 86 | + |
| 87 | +lemma one : map_fun f (1 : 𝕎 R) = 1 := by map_fun_tac |
| 88 | + |
| 89 | +lemma add : map_fun f (x + y) = map_fun f x + map_fun f y := by map_fun_tac |
| 90 | + |
| 91 | +lemma mul : map_fun f (x * y) = map_fun f x * map_fun f y := by map_fun_tac |
| 92 | + |
| 93 | +lemma neg : map_fun f (-x) = -map_fun f x := by map_fun_tac |
| 94 | + |
| 95 | +end map_fun |
| 96 | + |
| 97 | +end witt_vector |
| 98 | + |
| 99 | +section tactic |
| 100 | +setup_tactic_parser |
| 101 | +open tactic |
| 102 | + |
| 103 | +/-- An auxiliary tactic for proving that `ghost_fun` respects the ring operations. -/ |
| 104 | +meta def tactic.interactive.ghost_fun_tac (φ fn : parse parser.pexpr) : tactic unit := |
| 105 | +do fn ← to_expr ```(%%fn : fin _ → ℕ → R), |
| 106 | + `(fin %%k → _ → _) ← infer_type fn, |
| 107 | + `[ext n], |
| 108 | + to_expr ```(witt_structure_int_prop p (%%φ : mv_polynomial (fin %%k) ℤ) n) >>= note `aux none >>= |
| 109 | + apply_fun_to_hyp ```(aeval (uncurry %%fn)) none, |
| 110 | +`[simp only [aeval_bind₁] at aux, |
| 111 | + simp only [pi.zero_apply, pi.one_apply, pi.add_apply, pi.mul_apply, pi.neg_apply, ghost_fun], |
| 112 | + convert aux using 1; clear aux; |
| 113 | + simp only [alg_hom.map_zero, alg_hom.map_one, alg_hom.map_add, alg_hom.map_mul, alg_hom.map_neg, |
| 114 | + aeval_X, aeval_rename]; refl] |
| 115 | +end tactic |
| 116 | + |
| 117 | +namespace witt_vector |
| 118 | + |
| 119 | +/-- Evaluates the `n`th Witt polynomial on the first `n` coefficients of `x`, |
| 120 | +producing a value in `R`. |
| 121 | +This function will be bundled as the ring homomorphism `witt_vector.ghost_map` |
| 122 | +once the ring structure is available, |
| 123 | +but we rely on it to set up the ring structure in the first place. -/ |
| 124 | +private def ghost_fun : 𝕎 R → (ℕ → R) := λ x n, aeval x.coeff (W_ ℤ n) |
| 125 | + |
| 126 | +section ghost_fun |
| 127 | +include hp |
| 128 | + |
| 129 | +/- The following lemmas are not `@[simp]` because they will be bundled in `ghost_map` later on. -/ |
| 130 | + |
| 131 | +variables (x y : 𝕎 R) |
| 132 | + |
| 133 | +private lemma ghost_fun_zero : ghost_fun (0 : 𝕎 R) = 0 := by ghost_fun_tac 0 ![] |
| 134 | + |
| 135 | +private lemma ghost_fun_one : ghost_fun (1 : 𝕎 R) = 1 := by ghost_fun_tac 1 ![] |
| 136 | + |
| 137 | +private lemma ghost_fun_add : ghost_fun (x + y) = ghost_fun x + ghost_fun y := |
| 138 | +by ghost_fun_tac (X 0 + X 1) ![x.coeff, y.coeff] |
| 139 | + |
| 140 | +private lemma ghost_fun_mul : ghost_fun (x * y) = ghost_fun x * ghost_fun y := |
| 141 | +by ghost_fun_tac (X 0 * X 1) ![x.coeff, y.coeff] |
| 142 | + |
| 143 | +private lemma ghost_fun_neg : ghost_fun (-x) = - ghost_fun x := by ghost_fun_tac (-X 0) ![x.coeff] |
| 144 | + |
| 145 | +end ghost_fun |
| 146 | + |
| 147 | +variables (p) (R) |
| 148 | + |
| 149 | +/-- The bijection between `𝕎 R` and `ℕ → R`, under the assumption that `p` is invertible in `R`. |
| 150 | +In `witt_vector.ghost_equiv` we upgrade this to an isomorphism of rings. -/ |
| 151 | +private def ghost_equiv' [invertible (p : R)] : 𝕎 R ≃ (ℕ → R) := |
| 152 | +{ to_fun := ghost_fun, |
| 153 | + inv_fun := λ x, mk p $ λ n, aeval x (X_in_terms_of_W p R n), |
| 154 | + left_inv := |
| 155 | + begin |
| 156 | + intro x, |
| 157 | + ext n, |
| 158 | + have := bind₁_witt_polynomial_X_in_terms_of_W p R n, |
| 159 | + apply_fun (aeval x.coeff) at this, |
| 160 | + simpa only [aeval_bind₁, aeval_X, ghost_fun, aeval_witt_polynomial] |
| 161 | + end, |
| 162 | + right_inv := |
| 163 | + begin |
| 164 | + intro x, |
| 165 | + ext n, |
| 166 | + have := bind₁_X_in_terms_of_W_witt_polynomial p R n, |
| 167 | + apply_fun (aeval x) at this, |
| 168 | + simpa only [aeval_bind₁, aeval_X, ghost_fun, aeval_witt_polynomial] |
| 169 | + end } |
| 170 | + |
| 171 | +include hp |
| 172 | + |
| 173 | +local attribute [instance] |
| 174 | +private def comm_ring_aux₁ : comm_ring (𝕎 (mv_polynomial R ℚ)) := |
| 175 | +(ghost_equiv' p (mv_polynomial R ℚ)).injective.comm_ring (ghost_fun) |
| 176 | + ghost_fun_zero ghost_fun_one ghost_fun_add ghost_fun_mul ghost_fun_neg |
| 177 | + |
| 178 | +local attribute [instance] |
| 179 | +private def comm_ring_aux₂ : comm_ring (𝕎 (mv_polynomial R ℤ)) := |
| 180 | +(map_fun.injective _ $ map_injective (int.cast_ring_hom ℚ) int.cast_injective).comm_ring _ |
| 181 | + (map_fun.zero _) (map_fun.one _) (map_fun.add _) (map_fun.mul _) (map_fun.neg _) |
| 182 | + |
| 183 | +/-- The commutative ring structure on `𝕎 R`. -/ |
| 184 | +instance : comm_ring (𝕎 R) := |
| 185 | +(map_fun.surjective _ $ counit_surjective _).comm_ring (map_fun $ mv_polynomial.counit _) |
| 186 | + (map_fun.zero _) (map_fun.one _) (map_fun.add _) (map_fun.mul _) (map_fun.neg _) |
| 187 | + |
| 188 | +variables {p R} |
| 189 | + |
| 190 | +/-- `witt_vector.map f` is the ring homomorphism `𝕎 R →+* 𝕎 S` naturally induced |
| 191 | +by a ring homomorphism `f : R →+* S`. It acts coefficientwise. -/ |
| 192 | +def map (f : R →+* S) : 𝕎 R →+* 𝕎 S := |
| 193 | +{ to_fun := map_fun f, |
| 194 | + map_zero' := map_fun.zero f, |
| 195 | + map_one' := map_fun.one f, |
| 196 | + map_add' := map_fun.add f, |
| 197 | + map_mul' := map_fun.mul f } |
| 198 | + |
| 199 | +lemma map_injective (f : R →+* S) (hf : injective f) : injective (map f : 𝕎 R → 𝕎 S) := |
| 200 | +map_fun.injective f hf |
| 201 | + |
| 202 | +lemma map_surjective (f : R →+* S) (hf : surjective f) : surjective (map f : 𝕎 R → 𝕎 S) := |
| 203 | +map_fun.surjective f hf |
| 204 | + |
| 205 | +@[simp] lemma map_coeff (f : R →+* S) (x : 𝕎 R) (n : ℕ) : |
| 206 | + (map f x).coeff n = f (x.coeff n) := rfl |
| 207 | + |
| 208 | +/-- `witt_vector.ghost_map` is a ring homomorphism that maps each Witt vector |
| 209 | +to the sequence of its ghost components. -/ |
| 210 | +def ghost_map : 𝕎 R →+* ℕ → R := |
| 211 | +{ to_fun := ghost_fun, |
| 212 | + map_zero' := ghost_fun_zero, |
| 213 | + map_one' := ghost_fun_one, |
| 214 | + map_add' := ghost_fun_add, |
| 215 | + map_mul' := ghost_fun_mul } |
| 216 | + |
| 217 | +/-- Evaluates the `n`th Witt polynomial on the first `n` coefficients of `x`, |
| 218 | +producing a value in `R`. -/ |
| 219 | +def ghost_component (n : ℕ) : 𝕎 R →+* R := (ring_hom.apply _ n).comp ghost_map |
| 220 | + |
| 221 | +lemma ghost_component_apply (n : ℕ) (x : 𝕎 R) : ghost_component n x = aeval x.coeff (W_ ℤ n) := rfl |
| 222 | + |
| 223 | +@[simp] lemma ghost_map_apply (x : 𝕎 R) (n : ℕ) : ghost_map x n = ghost_component n x := rfl |
| 224 | + |
| 225 | +variables (p R) [invertible (p : R)] |
| 226 | + |
| 227 | +/-- `witt_vector.ghost_map` is a ring isomorphism when `p` is invertible in `R`. -/ |
| 228 | +def ghost_equiv : 𝕎 R ≃+* (ℕ → R) := |
| 229 | +{ .. (ghost_map : 𝕎 R →+* (ℕ → R)), .. (ghost_equiv' p R) } |
| 230 | + |
| 231 | +@[simp] lemma ghost_equiv_coe : (ghost_equiv p R : 𝕎 R →+* (ℕ → R)) = ghost_map := rfl |
| 232 | + |
| 233 | +lemma ghost_map.bijective_of_invertible : function.bijective (ghost_map : 𝕎 R → ℕ → R) := |
| 234 | +(ghost_equiv p R).bijective |
| 235 | + |
| 236 | +end witt_vector |
| 237 | + |
| 238 | +attribute [irreducible] witt_vector |
0 commit comments