docs: More doc strings
This commit is contained in:
parent
a00a1020a8
commit
c24029c9ca
7 changed files with 47 additions and 0 deletions
|
@ -48,8 +48,10 @@ We define the direct sum of the edge and vertex momentum spaces.
|
|||
Corresponding to that spanned by its momentum. -/
|
||||
def HalfEdgeMomenta : Type := F.𝓱𝓔 → ℝ
|
||||
|
||||
/-- The half momenta carries the structure of an addative commutative group. -/
|
||||
instance : AddCommGroup F.HalfEdgeMomenta := Pi.addCommGroup
|
||||
|
||||
/-- The half momenta carries the structure of a module over `ℝ`. Defined via its target. -/
|
||||
instance : Module ℝ F.HalfEdgeMomenta := Pi.module _ _ _
|
||||
|
||||
/-- An auxiliary function used to define the Euclidean inner product on `F.HalfEdgeMomenta`. -/
|
||||
|
@ -81,16 +83,20 @@ def euclidInner : F.HalfEdgeMomenta →ₗ[ℝ] F.HalfEdgeMomenta →ₗ[ℝ]
|
|||
Corresponding to that spanned by its total outflowing momentum. -/
|
||||
def EdgeMomenta : Type := F.𝓔 → ℝ
|
||||
|
||||
/-- The edge momenta form an addative commuative group. -/
|
||||
instance : AddCommGroup F.EdgeMomenta := Pi.addCommGroup
|
||||
|
||||
/-- The edge momenta form a module over `ℝ`. -/
|
||||
instance : Module ℝ F.EdgeMomenta := Pi.module _ _ _
|
||||
|
||||
/-- The type which assocaites to each ege a `1`-dimensional vector space.
|
||||
Corresponding to that spanned by its total inflowing momentum. -/
|
||||
def VertexMomenta : Type := F.𝓥 → ℝ
|
||||
|
||||
/-- The vertex momenta carries the structure of an addative commuative group. -/
|
||||
instance : AddCommGroup F.VertexMomenta := Pi.addCommGroup
|
||||
|
||||
/-- The vertex momenta carries the structure of a module over `ℝ`. -/
|
||||
instance : Module ℝ F.VertexMomenta := Pi.module _ _ _
|
||||
|
||||
/-- The map from `Fin 2` to `Type` landing on `EdgeMomenta` and `VertexMomenta`. -/
|
||||
|
@ -99,11 +105,15 @@ def EdgeVertexMomentaMap : Fin 2 → Type := fun i =>
|
|||
| 0 => F.EdgeMomenta
|
||||
| 1 => F.VertexMomenta
|
||||
|
||||
/-- The target of the map `EdgeVertexMomentaMap` is either the type of edge momenta
|
||||
or vertex momenta and thus carries the structure of an addative commuative group. -/
|
||||
instance (i : Fin 2) : AddCommGroup (EdgeVertexMomentaMap F i) :=
|
||||
match i with
|
||||
| 0 => instAddCommGroupEdgeMomenta F
|
||||
| 1 => instAddCommGroupVertexMomenta F
|
||||
|
||||
/-- The target of the map `EdgeVertexMomentaMap` is either the type of edge momenta
|
||||
or vertex momenta and thus carries the structure of a module over `ℝ`. -/
|
||||
instance (i : Fin 2) : Module ℝ (EdgeVertexMomentaMap F i) :=
|
||||
match i with
|
||||
| 0 => instModuleRealEdgeMomenta F
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue