diff options
author | Greg Brown <greg.brown@cl.cam.ac.uk> | 2022-04-02 11:41:51 +0100 |
---|---|---|
committer | Greg Brown <greg.brown@cl.cam.ac.uk> | 2022-04-02 11:59:21 +0100 |
commit | 2167866c53aa7f9cbb52e776bfb64f53acf3fa2c (patch) | |
tree | d9422bd08ee318b3fad90d03210f6a02a4c30783 /src/Helium/Algebra/Ordered/StrictTotal/Properties/AbelianGroup.agda | |
parent | 23e8afe152a84551491594aea133488523525410 (diff) |
Add more properties for ordered structures.
Diffstat (limited to 'src/Helium/Algebra/Ordered/StrictTotal/Properties/AbelianGroup.agda')
-rw-r--r-- | src/Helium/Algebra/Ordered/StrictTotal/Properties/AbelianGroup.agda | 60 |
1 files changed, 51 insertions, 9 deletions
diff --git a/src/Helium/Algebra/Ordered/StrictTotal/Properties/AbelianGroup.agda b/src/Helium/Algebra/Ordered/StrictTotal/Properties/AbelianGroup.agda index fe64b32..a6cbcab 100644 --- a/src/Helium/Algebra/Ordered/StrictTotal/Properties/AbelianGroup.agda +++ b/src/Helium/Algebra/Ordered/StrictTotal/Properties/AbelianGroup.agda @@ -14,18 +14,60 @@ module Helium.Algebra.Ordered.StrictTotal.Properties.AbelianGroup where open AbelianGroup abelianGroup - + renaming + ( trans to <-trans + ; irrefl to <-irrefl + ; asym to <-asym + ) open import Relation.Binary.Core open import Relation.Binary.Reasoning.StrictPartialOrder strictPartialOrder open import Algebra.Properties.AbelianGroup Unordered.abelianGroup public open import Algebra.Properties.CommutativeMonoid.Mult.TCOptimised Unordered.commutativeMonoid public +open import Helium.Relation.Binary.Properties.StrictTotalOrder strictTotalOrder public open import Helium.Algebra.Ordered.StrictTotal.Properties.Group group public - using (<⇒≱; ≤⇒≯; >⇒≉; ≈⇒≯; <⇒≉; ≈⇒≮; ≤∧≉⇒<; ≥∧≉⇒>; x<y⇒ε<yx⁻¹) - -⁻¹-anti-mono : _⁻¹ Preserves _<_ ⟶ _>_ -⁻¹-anti-mono {x} {y} x<y = begin-strict - y ⁻¹ ≈˘⟨ identityˡ (y ⁻¹) ⟩ - ε ∙ y ⁻¹ <⟨ ∙-invariantʳ (y ⁻¹) (x<y⇒ε<yx⁻¹ x<y) ⟩ - y ∙ x ⁻¹ ∙ y ⁻¹ ≈⟨ xyx⁻¹≈y y (x ⁻¹) ⟩ - x ⁻¹ ∎ + using + ( ∙-mono-<; ∙-monoˡ-<; ∙-monoʳ-< + ; ∙-mono-≤; ∙-monoˡ-≤; ∙-monoʳ-≤ + + ; ∙-cancelˡ-<; ∙-cancelʳ-<; ∙-cancel-< + ; ∙-cancelˡ-≤; ∙-cancelʳ-≤; ∙-cancel-≤ + -- _∙_ pres signs + ; x≥ε∧y>ε⇒x∙y>ε; x>ε∧y≥ε⇒x∙y>ε + ; x≤ε∧y<ε⇒x∙y<ε; x<ε∧y≤ε⇒x∙y<ε + ; x≥ε∧y≥ε⇒x∙y≥ε; x≤ε∧y≤ε⇒x∙y≤ε + -- _∙_ resp signs + ; x≤ε∧x∙y>ε⇒y>ε; x≤ε∧y∙x>ε⇒y>ε + ; x<ε∧x∙y≥ε⇒y>ε; x<ε∧y∙x≥ε⇒y>ε + ; x≥ε∧x∙y<ε⇒y<ε; x≥ε∧y∙x<ε⇒y<ε + ; x>ε∧x∙y≤ε⇒y<ε; x>ε∧y∙x≤ε⇒y<ε + ; x≤ε∧x∙y≥ε⇒y≥ε; x≤ε∧y∙x≥ε⇒y≥ε + ; x≥ε∧x∙y≤ε⇒y≤ε; x≥ε∧y∙x≤ε⇒y≤ε + + ; ×-zeroˡ; ×-zeroʳ + ; ×-identityˡ + + ; n≢0⇒×-monoˡ-<; x>ε⇒×-monoʳ-<; x<ε⇒×-anti-monoʳ-< + ; ×-monoˡ-≤; x≥ε⇒×-monoʳ-≤; x≤ε⇒×-anti-monoʳ-≤ + + ; ×-cancelˡ-<; x≥ε⇒×-cancelʳ-<; x≤ε⇒×-anti-cancelʳ-< + ; n≢0⇒×-cancelˡ-≤; x>ε⇒×-cancelʳ-≤; x<ε⇒×-anti-cancelʳ-≤ + -- _×_ pres signs + ; n≢0∧x>ε⇒n×x>ε; n≢0∧x<ε⇒n×x<ε + ; x≥ε⇒n×x≥ε; x≤ε⇒n×x≤ε + -- _×_ resp signs + ; n×x>ε⇒x>ε; n×x<ε⇒x<ε + ; n≢0∧n×x≥ε⇒x≥ε; n≢0∧n×x≤ε⇒x≤ε + + ; ⁻¹-anti-mono-<; ⁻¹-anti-mono-≤ + + ; ⁻¹-cancel; ⁻¹-anti-cancel-<; ⁻¹-anti-cancel-≤ + + ; x<ε⇒x⁻¹>ε; x>ε⇒x⁻¹<ε + ; x≤ε⇒x⁻¹≥ε; x≥ε⇒x⁻¹≤ε + + ; x⁻¹<ε⇒x>ε; x⁻¹>ε⇒x<ε + ; x⁻¹≤ε⇒x≥ε; x⁻¹≥ε⇒x≤ε + + ; x<y⇒ε<y∙x⁻¹ ; ε<y∙x⁻¹⇒x<y + ) |