summaryrefslogtreecommitdiff
path: root/src/Helium/Algebra/Ordered/StrictTotal/Properties/Group.agda
diff options
context:
space:
mode:
authorGreg Brown <greg.brown@cl.cam.ac.uk>2022-04-05 12:43:53 +0100
committerGreg Brown <greg.brown@cl.cam.ac.uk>2022-04-05 12:43:53 +0100
commit7de7db24ba5bb2cfe9fae6f070e60db3510fcc69 (patch)
tree50f3e5e49de31f640501844d8dd23f8baeba04a4 /src/Helium/Algebra/Ordered/StrictTotal/Properties/Group.agda
parent2f4a5a56437837cc90c9043131764513472347df (diff)
Add a useful transport proof.
Diffstat (limited to 'src/Helium/Algebra/Ordered/StrictTotal/Properties/Group.agda')
-rw-r--r--src/Helium/Algebra/Ordered/StrictTotal/Properties/Group.agda27
1 files changed, 27 insertions, 0 deletions
diff --git a/src/Helium/Algebra/Ordered/StrictTotal/Properties/Group.agda b/src/Helium/Algebra/Ordered/StrictTotal/Properties/Group.agda
index 81c57ad..cf2c9df 100644
--- a/src/Helium/Algebra/Ordered/StrictTotal/Properties/Group.agda
+++ b/src/Helium/Algebra/Ordered/StrictTotal/Properties/Group.agda
@@ -228,3 +228,30 @@ x≤y⇒ε≤y∙x⁻¹ {x} {y} x≤y = begin
y ∙ (x ⁻¹ ∙ x) ≈⟨ ∙-congˡ (inverseˡ x) ⟩
y ∙ ε ≈⟨ identityʳ y ⟩
y ∎
+
+---- Transport
+
+-- x∙y⁻¹<z⇒x<z∙y
+-- x⁻¹∙y<z⇒y<x∙z
+-- x∙y⁻¹>z⇒x>z∙y
+-- x⁻¹∙y>z⇒y>x∙z
+-- x∙y⁻¹≤z⇒x≤z∙y
+-- x⁻¹∙y≤z⇒y≤x∙z
+-- x∙y⁻¹≥z⇒x≥z∙y
+-- x⁻¹∙y≥z⇒y≥x∙z
+
+x<y∙z⇒x∙z⁻¹<y : ∀ {x y z} → x < y ∙ z → x ∙ z ⁻¹ < y
+x<y∙z⇒x∙z⁻¹<y {x} {y} {z} x<y∙z = ∙-cancelʳ-< _ _ (begin-strict
+ x ∙ z ⁻¹ ∙ z ≈⟨ assoc x _ z ⟩
+ x ∙ (z ⁻¹ ∙ z) ≈⟨ ∙-congˡ (inverseˡ z) ⟩
+ x ∙ ε ≈⟨ identityʳ x ⟩
+ x <⟨ x<y∙z ⟩
+ y ∙ z ∎)
+
+-- x<y∙z⇒y⁻¹∙x<z
+-- x>y∙z⇒x∙z⁻¹>y
+-- x>y∙z⇒y⁻¹∙x>z
+-- x≤y∙z⇒x∙z⁻¹≤y
+-- x≤y∙z⇒y⁻¹∙x≤z
+-- x≥y∙z⇒x∙z⁻¹≥y
+-- x≥y∙z⇒y⁻¹∙x≥z