|
| 1 | +/- |
| 2 | +Copyright (c) 2021 Eric Wieser. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Eric Wieser |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module algebra.char_p.two |
| 7 | +! leanprover-community/mathlib commit 7f1ba1a333d66eed531ecb4092493cd1b6715450 |
| 8 | +! Please do not edit these lines, except to modify the commit id |
| 9 | +! if you have ported upstream changes. |
| 10 | +-/ |
| 11 | +import Mathlib.Algebra.CharP.Basic |
| 12 | + |
| 13 | +/-! |
| 14 | +# Lemmas about rings of characteristic two |
| 15 | +
|
| 16 | +This file contains results about `CharP R 2`, in the `CharTwo` namespace. |
| 17 | +
|
| 18 | +The lemmas in this file with a `_sq` suffix are just special cases of the `_pow_char` lemmas |
| 19 | +elsewhere, with a shorter name for ease of discovery, and no need for a `[Fact (prime 2)]` argument. |
| 20 | +-/ |
| 21 | + |
| 22 | + |
| 23 | +variable {R ι : Type _} |
| 24 | + |
| 25 | +namespace CharTwo |
| 26 | + |
| 27 | +section Semiring |
| 28 | + |
| 29 | +variable [Semiring R] [CharP R 2] |
| 30 | + |
| 31 | +theorem two_eq_zero : (2 : R) = 0 := by rw [← Nat.cast_two, CharP.cast_eq_zero] |
| 32 | +#align char_two.two_eq_zero CharTwo.two_eq_zero |
| 33 | + |
| 34 | +@[simp] |
| 35 | +theorem add_self_eq_zero (x : R) : x + x = 0 := by rw [← two_smul R x, two_eq_zero, zero_smul] |
| 36 | +#align char_two.add_self_eq_zero CharTwo.add_self_eq_zero |
| 37 | + |
| 38 | +set_option linter.deprecated false in |
| 39 | +@[simp] |
| 40 | +theorem bit0_eq_zero : (bit0 : R → R) = 0 := by |
| 41 | + funext |
| 42 | + exact add_self_eq_zero _ |
| 43 | +#align char_two.bit0_eq_zero CharTwo.bit0_eq_zero |
| 44 | + |
| 45 | +set_option linter.deprecated false in |
| 46 | +theorem bit0_apply_eq_zero (x : R) : (bit0 x : R) = 0 := by simp |
| 47 | +#align char_two.bit0_apply_eq_zero CharTwo.bit0_apply_eq_zero |
| 48 | + |
| 49 | +set_option linter.deprecated false in |
| 50 | +@[simp] |
| 51 | +theorem bit1_eq_one : (bit1 : R → R) = 1 := by |
| 52 | + funext |
| 53 | + simp [bit1] |
| 54 | +#align char_two.bit1_eq_one CharTwo.bit1_eq_one |
| 55 | + |
| 56 | +set_option linter.deprecated false in |
| 57 | +theorem bit1_apply_eq_one (x : R) : (bit1 x : R) = 1 := by simp |
| 58 | +#align char_two.bit1_apply_eq_one CharTwo.bit1_apply_eq_one |
| 59 | + |
| 60 | +end Semiring |
| 61 | + |
| 62 | +section Ring |
| 63 | + |
| 64 | +variable [Ring R] [CharP R 2] |
| 65 | + |
| 66 | +@[simp] |
| 67 | +theorem neg_eq (x : R) : -x = x := by |
| 68 | + rw [neg_eq_iff_add_eq_zero, ← two_smul R x, two_eq_zero, zero_smul] |
| 69 | +#align char_two.neg_eq CharTwo.neg_eq |
| 70 | + |
| 71 | +theorem neg_eq' : Neg.neg = (id : R → R) := |
| 72 | + funext neg_eq |
| 73 | +#align char_two.neg_eq' CharTwo.neg_eq' |
| 74 | + |
| 75 | +@[simp] |
| 76 | +theorem sub_eq_add (x y : R) : x - y = x + y := by rw [sub_eq_add_neg, neg_eq] |
| 77 | +#align char_two.sub_eq_add CharTwo.sub_eq_add |
| 78 | + |
| 79 | +theorem sub_eq_add' : Sub.sub = ((· + ·) : R → R → R) := |
| 80 | + funext fun x => funext fun y => sub_eq_add x y |
| 81 | +#align char_two.sub_eq_add' CharTwo.sub_eq_add' |
| 82 | + |
| 83 | +end Ring |
| 84 | + |
| 85 | +section CommSemiring |
| 86 | + |
| 87 | +variable [CommSemiring R] [CharP R 2] |
| 88 | + |
| 89 | +theorem add_sq (x y : R) : (x + y) ^ 2 = x ^ 2 + y ^ 2 := |
| 90 | + add_pow_char _ _ _ |
| 91 | +#align char_two.add_sq CharTwo.add_sq |
| 92 | + |
| 93 | +theorem add_mul_self (x y : R) : (x + y) * (x + y) = x * x + y * y := by |
| 94 | + rw [← pow_two, ← pow_two, ← pow_two, add_sq] |
| 95 | +#align char_two.add_mul_self CharTwo.add_mul_self |
| 96 | + |
| 97 | +open BigOperators |
| 98 | + |
| 99 | +theorem list_sum_sq (l : List R) : l.sum ^ 2 = (l.map (· ^ 2)).sum := |
| 100 | + list_sum_pow_char _ _ |
| 101 | +#align char_two.list_sum_sq CharTwo.list_sum_sq |
| 102 | + |
| 103 | +theorem list_sum_mul_self (l : List R) : l.sum * l.sum = (List.map (fun x => x * x) l).sum := by |
| 104 | + simp_rw [← pow_two, list_sum_sq] |
| 105 | +#align char_two.list_sum_mul_self CharTwo.list_sum_mul_self |
| 106 | + |
| 107 | +theorem multiset_sum_sq (l : Multiset R) : l.sum ^ 2 = (l.map (· ^ 2)).sum := |
| 108 | + multiset_sum_pow_char _ _ |
| 109 | +#align char_two.multiset_sum_sq CharTwo.multiset_sum_sq |
| 110 | + |
| 111 | +theorem multiset_sum_mul_self (l : Multiset R) : |
| 112 | + l.sum * l.sum = (Multiset.map (fun x => x * x) l).sum := by simp_rw [← pow_two, multiset_sum_sq] |
| 113 | +#align char_two.multiset_sum_mul_self CharTwo.multiset_sum_mul_self |
| 114 | + |
| 115 | +theorem sum_sq (s : Finset ι) (f : ι → R) : (∑ i in s, f i) ^ 2 = ∑ i in s, f i ^ 2 := |
| 116 | + sum_pow_char _ _ _ |
| 117 | +#align char_two.sum_sq CharTwo.sum_sq |
| 118 | + |
| 119 | +theorem sum_mul_self (s : Finset ι) (f : ι → R) : |
| 120 | + ((∑ i in s, f i) * ∑ i in s, f i) = ∑ i in s, f i * f i := by simp_rw [← pow_two, sum_sq] |
| 121 | +#align char_two.sum_mul_self CharTwo.sum_mul_self |
| 122 | + |
| 123 | +end CommSemiring |
| 124 | + |
| 125 | +end CharTwo |
| 126 | + |
| 127 | +section ringChar |
| 128 | + |
| 129 | +variable [Ring R] |
| 130 | + |
| 131 | +theorem neg_one_eq_one_iff [Nontrivial R] : (-1 : R) = 1 ↔ ringChar R = 2 := by |
| 132 | + refine' ⟨fun h => _, fun h => @CharTwo.neg_eq _ _ (ringChar.of_eq h) 1⟩ |
| 133 | + rw [eq_comm, ← sub_eq_zero, sub_neg_eq_add, ← Nat.cast_one, ← Nat.cast_add] at h |
| 134 | + exact ((Nat.dvd_prime Nat.prime_two).mp (ringChar.dvd h)).resolve_left CharP.ringChar_ne_one |
| 135 | +#align neg_one_eq_one_iff neg_one_eq_one_iff |
| 136 | + |
| 137 | +@[simp] |
| 138 | +theorem orderOf_neg_one [Nontrivial R] : orderOf (-1 : R) = if ringChar R = 2 then 1 else 2 := by |
| 139 | + split_ifs with h |
| 140 | + · rw [neg_one_eq_one_iff.2 h, orderOf_one] |
| 141 | + apply orderOf_eq_prime |
| 142 | + · simp |
| 143 | + simpa [neg_one_eq_one_iff] using h |
| 144 | +#align order_of_neg_one orderOf_neg_one |
| 145 | + |
| 146 | +end ringChar |
0 commit comments