Lots of speedups (#116)

This commit is contained in:
Patrick Stevens
2020-04-16 13:41:51 +01:00
committed by GitHub
parent 1bcb3f8537
commit 9b80058157
63 changed files with 1082 additions and 564 deletions

View File

@@ -0,0 +1,19 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Orders.Total.Definition
open import Orders.Partial.Definition
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Setoids.Orders where
partialOrderToSetoidPartialOrder : {a b : _} {A : Set a} (P : PartialOrder {a} A {b}) SetoidPartialOrder (reflSetoid A) (PartialOrder._<_ P)
SetoidPartialOrder.<WellDefined (partialOrderToSetoidPartialOrder P) a=b c=d a<c rewrite a=b | c=d = a<c
SetoidPartialOrder.irreflexive (partialOrderToSetoidPartialOrder P) = PartialOrder.irreflexive P
SetoidPartialOrder.<Transitive (partialOrderToSetoidPartialOrder P) = PartialOrder.<Transitive P
totalOrderToSetoidTotalOrder : {a b : _} {A : Set a} (T : TotalOrder {a} A {b}) SetoidTotalOrder (partialOrderToSetoidPartialOrder (TotalOrder.order T))
SetoidTotalOrder.totality (totalOrderToSetoidTotalOrder T) = TotalOrder.totality T

View File

@@ -0,0 +1,25 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Orders.Total.Definition
open import Orders.Partial.Definition
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Setoids.Orders.Partial.Definition where
record SetoidPartialOrder {a b c : _} {A : Set a} (S : Setoid {a} {b} A) (_<_ : Rel {a} {c} A) : Set (a b c) where
open Setoid S
field
<WellDefined : {a b c d : A} (a b) (c d) a < c b < d
irreflexive : {x : A} (x < x) False
<Transitive : {a b c : A} (a < b) (b < c) (a < c)
_<=_ : Rel {a} {b c} A
a <= b = (a < b) || (a b)
partialOrderToSetoidPartialOrder : {a b : _} {A : Set a} (P : PartialOrder {a} A {b}) SetoidPartialOrder (reflSetoid A) (PartialOrder._<_ P)
SetoidPartialOrder.<WellDefined (partialOrderToSetoidPartialOrder P) a=b c=d a<c rewrite a=b | c=d = a<c
SetoidPartialOrder.irreflexive (partialOrderToSetoidPartialOrder P) = PartialOrder.irreflexive P
SetoidPartialOrder.<Transitive (partialOrderToSetoidPartialOrder P) = PartialOrder.<Transitive P

View File

@@ -0,0 +1,33 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Orders.Total.Definition
open import Orders.Partial.Definition
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Setoids.Orders.Total.Definition where
open import Setoids.Orders.Partial.Definition
record SetoidTotalOrder {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_<_ : Rel {a} {c} A} (P : SetoidPartialOrder S _<_) : Set (a b c) where
open Setoid S
field
totality : (a b : A) ((a < b) || (b < a)) || (a b)
partial : SetoidPartialOrder S _<_
partial = P
min : A A A
min a b with totality a b
min a b | inl (inl a<b) = a
min a b | inl (inr b<a) = b
min a b | inr a=b = a
max : A A A
max a b with totality a b
max a b | inl (inl a<b) = b
max a b | inl (inr b<a) = a
max a b | inr a=b = b
totalOrderToSetoidTotalOrder : {a b : _} {A : Set a} (T : TotalOrder {a} A {b}) SetoidTotalOrder (partialOrderToSetoidPartialOrder (TotalOrder.order T))
SetoidTotalOrder.totality (totalOrderToSetoidTotalOrder T) = TotalOrder.totality T

View File

@@ -0,0 +1,67 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Orders.Total.Definition
open import Orders.Partial.Definition
open import Setoids.Setoids
open import Setoids.Orders.Partial.Definition
open import Setoids.Orders.Total.Definition
open import Functions
open import Sets.EquivalenceRelations
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Setoids.Orders.Total.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {c : _} {_<_ : A A Set c} {P : SetoidPartialOrder S _<_} (T : SetoidTotalOrder P) where
open SetoidTotalOrder T
open SetoidPartialOrder P
open Setoid S
open Equivalence eq
maxInequalitiesR : {a b c : A} (a < b) (a < c) (a < max b c)
maxInequalitiesR {a} {b} {c} a<b a<c with totality b c
... | inl (inl x) = a<c
... | inl (inr x) = a<b
... | inr x = a<c
minInequalitiesR : {a b c : A} (a < b) (a < c) (a < min b c)
minInequalitiesR {a} {b} {c} a<b a<c with totality b c
... | inl (inl x) = a<b
... | inl (inr x) = a<c
... | inr x = a<b
maxInequalitiesL : {a b c : A} (a < c) (b < c) (max a b < c)
maxInequalitiesL {a} {b} {c} a<b a<c with totality a b
... | inl (inl x) = a<c
... | inl (inr x) = a<b
... | inr x = a<c
minInequalitiesL : {a b c : A} (a < c) (b < c) (min a b < c)
minInequalitiesL {a} {b} {c} a<b a<c with totality a b
... | inl (inl x) = a<b
... | inl (inr x) = a<c
... | inr x = a<b
minLessL : (a b : A) min a b <= a
minLessL a b with totality a b
... | inl (inl x) = inr reflexive
... | inl (inr x) = inl x
... | inr x = inr reflexive
minLessR : (a b : A) min a b <= b
minLessR a b with totality a b
... | inl (inl x) = inl x
... | inl (inr x) = inr reflexive
... | inr x = inr x
maxGreaterL : (a b : A) a <= max a b
maxGreaterL a b with totality a b
... | inl (inl x) = inl x
... | inl (inr x) = inr reflexive
... | inr x = inr x
maxGreaterR : (a b : A) b <= max a b
maxGreaterR a b with totality a b
... | inl (inl x) = inr reflexive
... | inl (inr x) = inl x
... | inr x = inr reflexive