feat: make informal_definition and informal_lemma commands (#300)
* make informal_definition and informal_lemma commands * drop the fields "math", "physics", and "proof" from InformalDefinition/InformalLemma and use docstrings instead * render informal docstring in dependency graph
This commit is contained in:
parent
6aab0ba3cd
commit
f8f94979ab
33 changed files with 666 additions and 1089 deletions
|
@ -241,27 +241,28 @@ def dualRepIsoDiscrete (c : S.C) : S.FD.obj (Discrete.mk c) ≅ S.FD.obj (Discre
|
|||
ext x
|
||||
exact S.fromDualRep_toDualRep_eq_self c x
|
||||
|
||||
/-- Given a `i : Fin n` the isomorphism between `S.F.obj (OverColor.mk c)` and
|
||||
`S.F.obj (OverColor.mk (Function.update c i (S.τ (c i))))` induced by `dualRepIsoDiscrete` acting on
|
||||
the `i`-th component of the color.
|
||||
-/
|
||||
informal_definition dualRepIso where
|
||||
math :≈ "Given a `i : Fin n` the isomorphism between `S.F.obj (OverColor.mk c)` and
|
||||
`S.F.obj (OverColor.mk (Function.update c i (S.τ (c i))))` induced by `dualRepIsoDiscrete`
|
||||
acting on the `i`-th component of the color."
|
||||
deps :≈ [``dualRepIsoDiscrete]
|
||||
deps := [``dualRepIsoDiscrete]
|
||||
|
||||
/-- Acting with `dualRepIso` on the fst component of a `unitTensor` returns a metric. -/
|
||||
informal_lemma dualRepIso_unitTensor_fst where
|
||||
math :≈ "Acting with `dualRepIso` on the fst component of a `unitTensor` returns a metric."
|
||||
deps :≈ [``dualRepIso, ``unitTensor, ``metricTensor]
|
||||
deps := [``dualRepIso, ``unitTensor, ``metricTensor]
|
||||
|
||||
/-- Acting with `dualRepIso` on the snd component of a `unitTensor` returns a metric. -/
|
||||
informal_lemma dualRepIso_unitTensor_snd where
|
||||
math :≈ "Acting with `dualRepIso` on the snd component of a `unitTensor` returns a metric."
|
||||
deps :≈ [``dualRepIso, ``unitTensor, ``metricTensor]
|
||||
deps := [``dualRepIso, ``unitTensor, ``metricTensor]
|
||||
|
||||
/-- Acting with `dualRepIso` on the fst component of a `metricTensor` returns a unitTensor. -/
|
||||
informal_lemma dualRepIso_metricTensor_fst where
|
||||
math :≈ "Acting with `dualRepIso` on the fst component of a `metricTensor` returns a unitTensor."
|
||||
deps :≈ [``dualRepIso, ``unitTensor, ``metricTensor]
|
||||
deps := [``dualRepIso, ``unitTensor, ``metricTensor]
|
||||
|
||||
/-- Acting with `dualRepIso` on the snd component of a `metricTensor` returns a unitTensor. -/
|
||||
informal_lemma dualRepIso_metricTensor_snd where
|
||||
math :≈ "Acting with `dualRepIso` on the snd component of a `metricTensor` returns a unitTensor."
|
||||
deps :≈ [``dualRepIso, ``unitTensor, ``metricTensor]
|
||||
deps := [``dualRepIso, ``unitTensor, ``metricTensor]
|
||||
|
||||
end TensorSpecies
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue