Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit f1531ea

Browse files
committed
feat(ring_theory/witt_vector): witt_sub, a demonstration of is_poly techniques (#5165)
Co-Authored-By: Rob Y. Lewis <rob.y.lewis@gmail.com>
1 parent f6273d4 commit f1531ea

File tree

1 file changed

+33
-0
lines changed

1 file changed

+33
-0
lines changed

src/ring_theory/witt_vector/is_poly.lean

Lines changed: 33 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -617,4 +617,37 @@ attribute [ghost_simps]
617617
nat.succ_ne_zero nat.add_sub_cancel nat.succ_eq_add_one
618618
if_true eq_self_iff_true if_false forall_true_iff forall_2_true_iff forall_3_true_iff
619619

620+
/-!
621+
### Subtraction of Witt vectors
622+
623+
In Lean, subtraction in a ring is by definition equal to `x + -y`.
624+
For Witt vectors, this means that subtraction is not defined in terms of
625+
the polynomials `witt_sub p`.
626+
627+
As a demonstration of some of the techniques developed in this file,
628+
we show by a computation that evaluating `witt_sub p` on the coefficients of `x` and `y`
629+
gives the coefficients of `x - y`.
630+
631+
For a more powerful demonstration, see `ring_theory/witt_vector/identities.lean`.
632+
-/
633+
634+
lemma sub_eq (x y : 𝕎 R) :
635+
x - y = eval (witt_sub p) ![x, y] :=
636+
begin
637+
apply is_poly₂.ext ((add_is_poly₂).comp_right
638+
(neg_is_poly)) ⟨witt_sub p, by intros; refl⟩ _ _ x y,
639+
unfreezingI { clear_dependent R }, introsI R _Rcr x y n,
640+
simp only [←sub_eq_add_neg, ring_hom.map_sub],
641+
symmetry,
642+
have := witt_structure_int_prop p (X 0 - X 1 : mv_polynomial (fin 2) ℤ) n,
643+
apply_fun (aeval (function.uncurry ![x.coeff, y.coeff])) at this,
644+
simp only [aeval_bind₁, alg_hom.map_sub, bind₁_X_right] at this,
645+
simp only [aeval_eq_eval₂_hom, eval₂_hom_rename] at this,
646+
exact this,
647+
end
648+
649+
lemma sub_coeff (x y : 𝕎 R) (n : ℕ) :
650+
(x - y).coeff n = peval (witt_sub p n) ![x.coeff, y.coeff] :=
651+
by { rw [sub_eq], refl }
652+
620653
end witt_vector

0 commit comments

Comments
 (0)