mirror of
https://github.com/SinTan1729/lean-talk-sp24.git
synced 2024-12-24 21:08:36 -06:00
new: Added some algebraic structures
This commit is contained in:
parent
522424f97f
commit
a91eef1132
2 changed files with 43 additions and 0 deletions
|
@ -4,3 +4,5 @@ import «LeanTalkSP24».basics
|
|||
import «LeanTalkSP24».basics2
|
||||
import «LeanTalkSP24».infinitely_many_primes
|
||||
import «LeanTalkSP24».polynomial_over_field_dim_one
|
||||
import «LeanTalkSP24».alg_structures
|
||||
|
||||
|
|
41
LeanTalkSP24/alg_structures.lean
Normal file
41
LeanTalkSP24/alg_structures.lean
Normal file
|
@ -0,0 +1,41 @@
|
|||
import Mathlib.Data.Real.Basic
|
||||
|
||||
class Group₁ (α : Type _) where
|
||||
mul : α → α → α
|
||||
one : α
|
||||
inv : α → α
|
||||
mul_assoc : ∀ x y z : α, mul (mul x y) z = mul x (mul y z)
|
||||
mul_one : ∀ x : α, mul x one = x
|
||||
one_mul : ∀ x : α, mul one x = x
|
||||
mul_left_inv : ∀ x : α, mul (inv x) x = one
|
||||
|
||||
@[ext]
|
||||
structure Point where
|
||||
x : ℝ
|
||||
y : ℝ
|
||||
z : ℝ
|
||||
|
||||
namespace Point
|
||||
|
||||
def add (a b : Point) : Point :=
|
||||
⟨a.x + b.x, a.y + b.y, a.z + b.z⟩
|
||||
|
||||
def neg (a : Point) : Point := ⟨-a.x, -a.y, -a.z⟩
|
||||
|
||||
def zero : Point := ⟨0,0,0⟩
|
||||
|
||||
instance add_group_point: Group₁ Point where
|
||||
mul := add
|
||||
one := zero
|
||||
inv := neg
|
||||
mul_assoc := by simp [add, add_assoc]
|
||||
mul_one := by simp [add, zero]
|
||||
one_mul := by simp [add, zero]
|
||||
mul_left_inv := by simp [add, neg, zero]
|
||||
|
||||
end Point
|
||||
|
||||
instance hasMulGroup₁ {α : Type _} [Group₁ α] : Add α := ⟨Group₁.mul⟩
|
||||
|
||||
def v : Point := ⟨1,1,1⟩
|
||||
#check v+v
|
Loading…
Reference in a new issue