Symmetrized algebra #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
A commutative multiplication on a real or complex space can be constructed from any multiplication by "symmetrization" i.e $$ a \circ b = \frac{1}{2}(ab + ba) $$
We provide the symmetrized version of a type α as sym_alg α, with notation αˢʸᵐ.
Implementation notes #
The approach taken here is inspired by algebra.opposites. We use Oxford Spellings (IETF en-GB-oxendict).
References #
The symmetrized algebra has the same underlying space as the original algebra.
Instances for sym_alg
        
        - sym_alg.nontrivial
- sym_alg.inhabited
- sym_alg.subsingleton
- sym_alg.unique
- sym_alg.is_empty
- sym_alg.has_one
- sym_alg.has_zero
- sym_alg.has_add
- sym_alg.has_sub
- sym_alg.has_neg
- sym_alg.has_mul
- sym_alg.has_inv
- sym_alg.has_smul
- sym_alg.add_comm_semigroup
- sym_alg.add_monoid
- sym_alg.add_group
- sym_alg.add_comm_monoid
- sym_alg.add_comm_group
- sym_alg.module
- sym_alg.non_assoc_semiring
- sym_alg.non_assoc_ring
- sym_alg.is_comm_jordan
The element of sym_alg α that represents a : α.
Equations
The element of α represented by x : αˢʸᵐ.
Equations
Equations
Equations
Equations
- sym_alg.has_zero = {zero := ⇑sym_alg.sym 0}
Equations
- sym_alg.has_one = {one := ⇑sym_alg.sym 1}
Equations
- sym_alg.has_add = {add := λ (a b : αˢʸᵐ), ⇑sym_alg.sym (⇑sym_alg.unsym a + ⇑sym_alg.unsym b)}
Equations
- sym_alg.has_sub = {sub := λ (a b : αˢʸᵐ), ⇑sym_alg.sym (⇑sym_alg.unsym a - ⇑sym_alg.unsym b)}
Equations
- sym_alg.has_neg = {neg := λ (a : αˢʸᵐ), ⇑sym_alg.sym (-⇑sym_alg.unsym a)}
Equations
- sym_alg.has_mul = {mul := λ (a b : αˢʸᵐ), ⇑sym_alg.sym (⅟ 2 * (⇑sym_alg.unsym a * ⇑sym_alg.unsym b + ⇑sym_alg.unsym b * ⇑sym_alg.unsym a))}
Equations
- sym_alg.has_inv = {inv := λ (a : αˢʸᵐ), ⇑sym_alg.sym (⇑sym_alg.unsym a)⁻¹}
Equations
- sym_alg.has_smul R = {smul := λ (r : R) (a : αˢʸᵐ), ⇑sym_alg.sym (r • ⇑sym_alg.unsym a)}
Equations
- sym_alg.add_comm_semigroup = function.injective.add_comm_semigroup ⇑sym_alg.unsym sym_alg.unsym_injective sym_alg.add_comm_semigroup._proof_1
Equations
- sym_alg.add_monoid = function.injective.add_monoid ⇑sym_alg.unsym sym_alg.unsym_injective sym_alg.add_monoid._proof_1 sym_alg.add_monoid._proof_2 sym_alg.add_monoid._proof_3
Equations
- sym_alg.add_group = function.injective.add_group ⇑sym_alg.unsym sym_alg.unsym_injective sym_alg.add_group._proof_1 sym_alg.add_group._proof_2 sym_alg.add_group._proof_3 sym_alg.add_group._proof_4 sym_alg.add_group._proof_5 sym_alg.add_group._proof_6
Equations
- sym_alg.add_comm_monoid = {add := add_comm_semigroup.add sym_alg.add_comm_semigroup, add_assoc := _, zero := add_monoid.zero sym_alg.add_monoid, zero_add := _, add_zero := _, nsmul := add_monoid.nsmul sym_alg.add_monoid, nsmul_zero' := _, nsmul_succ' := _, add_comm := _}
Equations
- sym_alg.add_comm_group = {add := add_comm_monoid.add sym_alg.add_comm_monoid, add_assoc := _, zero := add_comm_monoid.zero sym_alg.add_comm_monoid, zero_add := _, add_zero := _, nsmul := add_comm_monoid.nsmul sym_alg.add_comm_monoid, nsmul_zero' := _, nsmul_succ' := _, neg := add_group.neg sym_alg.add_group, sub := add_group.sub sym_alg.add_group, sub_eq_add_neg := _, zsmul := add_group.zsmul sym_alg.add_group, zsmul_zero' := _, zsmul_succ' := _, zsmul_neg' := _, add_left_neg := _, add_comm := _}
Equations
- sym_alg.module = function.injective.module R {to_fun := ⇑sym_alg.unsym, map_zero' := _, map_add' := _} sym_alg.unsym_injective sym_alg.module._proof_3
Equations
- sym_alg.sym.invertible a = {inv_of := ⇑sym_alg.sym (⅟ a), inv_of_mul_self := _, mul_inv_of_self := _}
Equations
- sym_alg.non_assoc_semiring = {add := add_comm_monoid.add sym_alg.add_comm_monoid, add_assoc := _, zero := 0, zero_add := _, add_zero := _, nsmul := add_comm_monoid.nsmul sym_alg.add_comm_monoid, nsmul_zero' := _, nsmul_succ' := _, add_comm := _, mul := has_mul.mul sym_alg.has_mul, left_distrib := _, right_distrib := _, zero_mul := _, mul_zero := _, one := 1, one_mul := _, mul_one := _, nat_cast := add_comm_monoid_with_one.nat_cast._default 1 add_comm_monoid.add sym_alg.non_assoc_semiring._proof_13 0 sym_alg.non_assoc_semiring._proof_14 sym_alg.non_assoc_semiring._proof_15 add_comm_monoid.nsmul sym_alg.non_assoc_semiring._proof_16 sym_alg.non_assoc_semiring._proof_17, nat_cast_zero := _, nat_cast_succ := _}
The symmetrization of a real (unital, associative) algebra is a non-associative ring.
Equations
- sym_alg.non_assoc_ring = {add := non_assoc_semiring.add sym_alg.non_assoc_semiring, add_assoc := _, zero := non_assoc_semiring.zero sym_alg.non_assoc_semiring, zero_add := _, add_zero := _, nsmul := non_assoc_semiring.nsmul sym_alg.non_assoc_semiring, nsmul_zero' := _, nsmul_succ' := _, neg := add_comm_group.neg sym_alg.add_comm_group, sub := add_comm_group.sub sym_alg.add_comm_group, sub_eq_add_neg := _, zsmul := add_comm_group.zsmul sym_alg.add_comm_group, zsmul_zero' := _, zsmul_succ' := _, zsmul_neg' := _, add_left_neg := _, add_comm := _, mul := non_assoc_semiring.mul sym_alg.non_assoc_semiring, left_distrib := _, right_distrib := _, zero_mul := _, mul_zero := _, one := non_assoc_semiring.one sym_alg.non_assoc_semiring, one_mul := _, mul_one := _, nat_cast := non_assoc_semiring.nat_cast sym_alg.non_assoc_semiring, nat_cast_zero := _, nat_cast_succ := _, int_cast := add_comm_group_with_one.int_cast._default non_assoc_semiring.nat_cast non_assoc_semiring.add sym_alg.non_assoc_ring._proof_20 non_assoc_semiring.zero sym_alg.non_assoc_ring._proof_21 sym_alg.non_assoc_ring._proof_22 non_assoc_semiring.nsmul sym_alg.non_assoc_ring._proof_23 sym_alg.non_assoc_ring._proof_24 non_assoc_semiring.one sym_alg.non_assoc_ring._proof_25 sym_alg.non_assoc_ring._proof_26 add_comm_group.neg add_comm_group.sub sym_alg.non_assoc_ring._proof_27 add_comm_group.zsmul sym_alg.non_assoc_ring._proof_28 sym_alg.non_assoc_ring._proof_29 sym_alg.non_assoc_ring._proof_30 sym_alg.non_assoc_ring._proof_31, int_cast_of_nat := _, int_cast_neg_succ_of_nat := _}
The squaring operation coincides for both multiplications
Equations
- sym_alg.is_comm_jordan = {mul_comm := _, lmul_comm_rmul_rmul := _}