diff options
author | Chloe Brown <chloe.brown.00@outlook.com> | 2021-03-29 18:08:09 +0100 |
---|---|---|
committer | Chloe Brown <chloe.brown.00@outlook.com> | 2021-03-29 18:08:09 +0100 |
commit | 13e0839831a528d26478a3a94c7470204460cce4 (patch) | |
tree | b907e2dee7ef170c879d3ec182bcc5b5eff374dd /src/Cfe/Language/Properties.agda | |
parent | da35892c463cd6b9a492c6aee09726a41031ca93 (diff) |
Introduce < for Languages.
Move around some definitions.
Diffstat (limited to 'src/Cfe/Language/Properties.agda')
-rw-r--r-- | src/Cfe/Language/Properties.agda | 25 |
1 files changed, 23 insertions, 2 deletions
diff --git a/src/Cfe/Language/Properties.agda b/src/Cfe/Language/Properties.agda index 756877c..52b5470 100644 --- a/src/Cfe/Language/Properties.agda +++ b/src/Cfe/Language/Properties.agda @@ -8,8 +8,6 @@ module Cfe.Language.Properties open Setoid over using () renaming (Carrier to C) open import Cfe.Language.Base over --- open Language - open import Data.List open import Data.List.Relation.Binary.Equality.Setoid over open import Function @@ -95,6 +93,29 @@ setoid a = record { isEquivalence = ≈-isEquivalence {a} } poset : ∀ a → Poset (c ⊔ ℓ ⊔ suc a) (c ⊔ ℓ ⊔ a) (c ⊔ a) poset a = record { isPartialOrder = ≤-isPartialOrder a } +<-trans : ∀ {a b c} → Trans (_<_ {a}) (_<_ {b} {c}) _<_ +<-trans A<B B<C = record + { f = B<C.f ∘ A<B.f + ; l = A<B.l + ; l∉A = A<B.l∉A + ; l∈B = B<C.f A<B.l∈B + } + where + module A<B = _<_ A<B + module B<C = _<_ B<C + +<-irrefl : ∀ {a b} → Irreflexive (_≈_ {a} {b}) _<_ +<-irrefl A≈B A<B = A<B.l∉A (A≈B.f⁻¹ A<B.l∈B) + where + module A≈B = _≈_ A≈B + module A<B = _<_ A<B + +<-asym : ∀ {a} → Asymmetric (_<_ {a}) +<-asym A<B B<A = A<B.l∉A (B<A.f A<B.l∈B) + where + module A<B = _<_ A<B + module B<A = _<_ B<A + lift-cong : ∀ {a} b (L : Language a) → Lift b L ≈ L lift-cong b L = record { f = lower |