|
| 1 | +/- |
| 2 | +Copyright (c) 2023 Eric Wieser. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Eric Wieser |
| 5 | +-/ |
| 6 | + |
| 7 | +import linear_algebra.quadratic_form.isometry |
| 8 | +import linear_algebra.quadratic_form.prod |
| 9 | +/-! |
| 10 | +# Quadratic form structures related to `module.dual` |
| 11 | +
|
| 12 | +## Main definitions |
| 13 | +
|
| 14 | +* `bilin_form.dual_prod R M`, the bilinear form on `(f, x) : module.dual R M × M` defined as |
| 15 | + `f x`. |
| 16 | +* `quadratic_form.dual_prod R M`, the quadratic form on `(f, x) : module.dual R M × M` defined as |
| 17 | + `f x`. |
| 18 | +* `quadratic_form.to_dual_prod : M × M →ₗ[R] module.dual R M × M` a form-preserving map from |
| 19 | + `(Q.prod $ -Q)` to `quadratic_form.dual_prod R M`. Note that we do not have the morphism |
| 20 | + version of `quadratic_form.isometry`, so for now this is stated without full bundling. |
| 21 | +
|
| 22 | +-/ |
| 23 | + |
| 24 | +variables (R M N : Type*) |
| 25 | + |
| 26 | +namespace bilin_form |
| 27 | + |
| 28 | +section semiring |
| 29 | +variables [comm_semiring R] [add_comm_monoid M] [module R M] |
| 30 | + |
| 31 | +/-- The symmetric bilinear form on `module.dual R M × M` defined as |
| 32 | +`B (f, x) (g, y) = f y + g x`. -/ |
| 33 | +@[simps] |
| 34 | +def dual_prod : bilin_form R (module.dual R M × M) := |
| 35 | +linear_map.to_bilin $ |
| 36 | + (linear_map.applyₗ.comp (linear_map.snd R (module.dual R M) M)).compl₂ |
| 37 | + (linear_map.fst R (module.dual R M) M) + |
| 38 | + ((linear_map.applyₗ.comp (linear_map.snd R (module.dual R M) M)).compl₂ |
| 39 | + (linear_map.fst R (module.dual R M) M)).flip |
| 40 | + |
| 41 | +lemma is_symm_dual_prod : (dual_prod R M).is_symm := |
| 42 | +λ x y, add_comm _ _ |
| 43 | + |
| 44 | +end semiring |
| 45 | + |
| 46 | +section ring |
| 47 | +variables [comm_ring R] [add_comm_group M] [module R M] |
| 48 | + |
| 49 | +lemma nondenerate_dual_prod : |
| 50 | + (dual_prod R M).nondegenerate ↔ function.injective (module.dual.eval R M) := |
| 51 | +begin |
| 52 | + classical, |
| 53 | + rw nondegenerate_iff_ker_eq_bot, |
| 54 | + rw linear_map.ker_eq_bot, |
| 55 | + let e := linear_equiv.prod_comm R _ _ |
| 56 | + ≪≫ₗ module.dual_prod_dual_equiv_dual R (module.dual R M) M, |
| 57 | + let h_d := e.symm.to_linear_map.comp (dual_prod R M).to_lin, |
| 58 | + refine (function.injective.of_comp_iff e.symm.injective (dual_prod R M).to_lin).symm.trans _, |
| 59 | + rw [←linear_equiv.coe_to_linear_map, ←linear_map.coe_comp], |
| 60 | + change function.injective h_d ↔ _, |
| 61 | + have : h_d = linear_map.prod_map (linear_map.id) (module.dual.eval R M), |
| 62 | + { refine linear_map.ext (λ x, prod.ext _ _), |
| 63 | + { ext, |
| 64 | + dsimp [h_d, module.dual.eval, linear_equiv.prod_comm], |
| 65 | + simp }, |
| 66 | + { ext, |
| 67 | + dsimp [h_d, module.dual.eval, linear_equiv.prod_comm], |
| 68 | + simp } }, |
| 69 | + rw [this, linear_map.coe_prod_map], |
| 70 | + refine prod.map_injective.trans _, |
| 71 | + exact and_iff_right function.injective_id, |
| 72 | +end |
| 73 | + |
| 74 | +end ring |
| 75 | + |
| 76 | +end bilin_form |
| 77 | + |
| 78 | +namespace quadratic_form |
| 79 | + |
| 80 | +section semiring |
| 81 | +variables [comm_semiring R] [add_comm_monoid M] [add_comm_monoid N] [module R M] [module R N] |
| 82 | + |
| 83 | +/-- The quadratic form on `module.dual R M × M` defined as `Q (f, x) = f x`. -/ |
| 84 | +@[simps] |
| 85 | +def dual_prod : quadratic_form R (module.dual R M × M) := |
| 86 | +{ to_fun := λ p, p.1 p.2, |
| 87 | + to_fun_smul := λ a p, by rw [prod.smul_fst, prod.smul_snd, linear_map.smul_apply, |
| 88 | + linear_map.map_smul, smul_eq_mul, smul_eq_mul, mul_assoc], |
| 89 | + exists_companion' := ⟨bilin_form.dual_prod R M, λ p q, begin |
| 90 | + rw [bilin_form.dual_prod_apply, prod.fst_add, prod.snd_add, linear_map.add_apply, map_add, |
| 91 | + map_add, add_right_comm _ (q.1 q.2), add_comm (q.1 p.2) (p.1 q.2), ←add_assoc, ←add_assoc], |
| 92 | + end⟩ } |
| 93 | + |
| 94 | +@[simp] |
| 95 | +lemma _root_.bilin_form.dual_prod.to_quadratic_form : |
| 96 | + (bilin_form.dual_prod R M).to_quadratic_form = 2 • dual_prod R M := |
| 97 | +ext $ λ a, (two_nsmul _).symm |
| 98 | + |
| 99 | +variables {R M N} |
| 100 | + |
| 101 | +/-- Any module isomorphism induces a quadratic isomorphism between the corresponding `dual_prod.` -/ |
| 102 | +@[simps] |
| 103 | +def dual_prod_isometry (f : M ≃ₗ[R] N) : |
| 104 | + (dual_prod R M).isometry (dual_prod R N) := |
| 105 | +{ to_linear_equiv := f.dual_map.symm.prod f, |
| 106 | + map_app' := λ x, fun_like.congr_arg x.fst $ f.symm_apply_apply _ } |
| 107 | + |
| 108 | +/-- `quadratic_form.dual_prod` commutes (isometrically) with `quadratic_form.prod`. -/ |
| 109 | +@[simps] |
| 110 | +def dual_prod_prod_isometry : |
| 111 | + (dual_prod R (M × N)).isometry ((dual_prod R M).prod (dual_prod R N)) := |
| 112 | +{ to_linear_equiv := |
| 113 | + ((module.dual_prod_dual_equiv_dual R M N).symm.prod (linear_equiv.refl R (M × N))) |
| 114 | + ≪≫ₗ linear_equiv.prod_prod_prod_comm R _ _ M N, |
| 115 | + map_app' := λ m, |
| 116 | + (m.fst.map_add _ _).symm.trans $ fun_like.congr_arg m.fst $ prod.ext (add_zero _) (zero_add _) } |
| 117 | + |
| 118 | +end semiring |
| 119 | + |
| 120 | +section ring |
| 121 | +variables [comm_ring R] [add_comm_group M] [module R M] |
| 122 | + |
| 123 | +variables {R M} |
| 124 | + |
| 125 | +/-- The isometry sending `(Q.prod $ -Q)` to `(quadratic_form.dual_prod R M)`. |
| 126 | +
|
| 127 | +This is `σ` from Proposition 4.8, page 84 of |
| 128 | +[*Hermitian K-Theory and Geometric Applications*][hyman1973]; though we swap the order of the pairs. |
| 129 | +-/ |
| 130 | +@[simps] |
| 131 | +def to_dual_prod (Q : quadratic_form R M) [invertible (2 : R)] : |
| 132 | + M × M →ₗ[R] module.dual R M × M := |
| 133 | +linear_map.prod |
| 134 | + (Q.associated.to_lin.comp (linear_map.fst _ _ _) |
| 135 | + + Q.associated.to_lin.comp (linear_map.snd _ _ _)) |
| 136 | + ((linear_map.fst _ _ _ - linear_map.snd _ _ _)) |
| 137 | + |
| 138 | +lemma to_dual_prod_isometry [invertible (2 : R)] (Q : quadratic_form R M) (x : M × M) : |
| 139 | + quadratic_form.dual_prod R M (to_dual_prod Q x) = (Q.prod $ -Q) x := |
| 140 | +begin |
| 141 | + dsimp only [to_dual_prod, associated, associated_hom], |
| 142 | + dsimp, |
| 143 | + simp [polar_comm _ x.1 x.2, ←sub_add, mul_sub, sub_mul, smul_sub, submonoid.smul_def, |
| 144 | + ←sub_eq_add_neg (Q x.1) (Q x.2)], |
| 145 | +end |
| 146 | + |
| 147 | +-- TODO: show that `to_dual_prod` is an equivalence |
| 148 | + |
| 149 | +end ring |
| 150 | + |
| 151 | +end quadratic_form |
0 commit comments