mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-13 23:58:38 +00:00
Reshuffle orders (#91)
This commit is contained in:
16
Orders/Partial/Definition.agda
Normal file
16
Orders/Partial/Definition.agda
Normal file
@@ -0,0 +1,16 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
|
||||
module Orders.Partial.Definition {a : _} (carrier : Set a) where
|
||||
|
||||
record PartialOrder {b : _} : Set (a ⊔ lsuc b) where
|
||||
field
|
||||
_<_ : Rel {a} {b} carrier
|
||||
irreflexive : {x : carrier} → (x < x) → False
|
||||
<Transitive : {a b c : carrier} → (a < b) → (b < c) → (a < c)
|
||||
<WellDefined : {r s t u : carrier} → (r ≡ t) → (s ≡ u) → r < s → t < u
|
||||
<WellDefined refl refl r<s = r<s
|
36
Orders/Total/Definition.agda
Normal file
36
Orders/Total/Definition.agda
Normal file
@@ -0,0 +1,36 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
|
||||
module Orders.Total.Definition {a : _} (carrier : Set a) where
|
||||
|
||||
open import Orders.Partial.Definition carrier
|
||||
|
||||
record TotalOrder {b : _} : Set (a ⊔ lsuc b) where
|
||||
field
|
||||
order : PartialOrder {b}
|
||||
_<_ : Rel carrier
|
||||
_<_ = PartialOrder._<_ order
|
||||
_≤_ : Rel carrier
|
||||
_≤_ a b = (a < b) || (a ≡ b)
|
||||
field
|
||||
totality : (a b : carrier) → ((a < b) || (b < a)) || (a ≡ b)
|
||||
|
||||
min : carrier → carrier → carrier
|
||||
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 : carrier → carrier → carrier
|
||||
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
|
||||
|
||||
irreflexive = PartialOrder.irreflexive order
|
||||
<Transitive = PartialOrder.<Transitive order
|
||||
<WellDefined = PartialOrder.<WellDefined order
|
||||
|
81
Orders/Total/Lemmas.agda
Normal file
81
Orders/Total/Lemmas.agda
Normal file
@@ -0,0 +1,81 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
|
||||
open import Orders.Total.Definition
|
||||
|
||||
module Orders.Total.Lemmas {a b : _} {A : Set a} (order : TotalOrder A {b}) where
|
||||
|
||||
open TotalOrder order
|
||||
|
||||
equivMin : {x y : A} → (x < y) → min x y ≡ x
|
||||
equivMin {x} {y} x<y with totality x y
|
||||
equivMin {x} {y} x<y | inl (inl x₁) = refl
|
||||
equivMin {x} {y} x<y | inl (inr y<x) = exFalso (irreflexive (<Transitive x<y y<x))
|
||||
equivMin {x} {y} x<y | inr x=y rewrite x=y = refl
|
||||
|
||||
equivMin' : {x y : A} → (min x y ≡ x) → (x < y) || (x ≡ y)
|
||||
equivMin' {x} {y} minEq with totality x y
|
||||
equivMin' {x} {y} minEq | inl (inl x<y) = inl x<y
|
||||
equivMin' {x} {y} minEq | inl (inr y<x) = exFalso (irreflexive (identityOfIndiscernablesLeft _<_ y<x minEq))
|
||||
equivMin' {x} {y} minEq | inr x=y = inr x=y
|
||||
|
||||
minCommutes : (x y : A) → (min x y) ≡ (min y x)
|
||||
minCommutes x y with totality x y
|
||||
minCommutes x y | inl (inl x<y) with totality y x
|
||||
minCommutes x y | inl (inl x<y) | inl (inl y<x) = exFalso (irreflexive (<Transitive y<x x<y))
|
||||
minCommutes x y | inl (inl x<y) | inl (inr x<y') = refl
|
||||
minCommutes x y | inl (inl x<y) | inr y=x = equalityCommutative y=x
|
||||
minCommutes x y | inl (inr y<x) with totality y x
|
||||
minCommutes x y | inl (inr y<x) | inl (inl y<x') = refl
|
||||
minCommutes x y | inl (inr y<x) | inl (inr x<y) = exFalso (irreflexive (<Transitive y<x x<y))
|
||||
minCommutes x y | inl (inr y<x) | inr y=x = refl
|
||||
minCommutes x y | inr x=y with totality y x
|
||||
minCommutes x y | inr x=y | inl (inl x₁) = x=y
|
||||
minCommutes x y | inr x=y | inl (inr x₁) = refl
|
||||
minCommutes x y | inr x=y | inr x₁ = x=y
|
||||
|
||||
minIdempotent : (x : A) → min x x ≡ x
|
||||
minIdempotent x with totality x x
|
||||
minIdempotent x | inl (inl x₁) = refl
|
||||
minIdempotent x | inl (inr x₁) = refl
|
||||
minIdempotent x | inr x₁ = refl
|
||||
|
||||
swapMin : {x y z : A} → (min x (min y z)) ≡ min y (min x z)
|
||||
swapMin {x} {y} {z} with totality y z
|
||||
swapMin {x} {y} {z} | inl (inl y<z) with totality x z
|
||||
swapMin {x} {y} {z} | inl (inl y<z) | inl (inl x<z) = minCommutes x y
|
||||
swapMin {x} {y} {z} | inl (inl y<z) | inl (inr z<x) with totality x y
|
||||
swapMin {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inl (inl x<y) = exFalso (irreflexive (<Transitive y<z (<Transitive z<x x<y)))
|
||||
swapMin {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inl (inr y<x) = equalityCommutative (equivMin y<z)
|
||||
swapMin {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inr x=y rewrite x=y = equalityCommutative (equivMin y<z)
|
||||
swapMin {x} {y} {z} | inl (inl y<z) | inr x=z = minCommutes x y
|
||||
swapMin {x} {y} {z} | inl (inr z<y) with totality x z
|
||||
swapMin {x} {y} {z} | inl (inr z<y) | inl (inl x<z) rewrite minCommutes y x = equalityCommutative (equivMin (<Transitive x<z z<y))
|
||||
swapMin {x} {y} {z} | inl (inr z<y) | inl (inr z<x) with totality y z
|
||||
swapMin {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inl (inl y<z) = exFalso (irreflexive (<Transitive z<y y<z))
|
||||
swapMin {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inl (inr z<y') = refl
|
||||
swapMin {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inr y=z = equalityCommutative y=z
|
||||
swapMin {x} {y} {z} | inl (inr z<y) | inr x=z rewrite x=z | minCommutes y z = equalityCommutative (equivMin z<y)
|
||||
swapMin {x} {y} {z} | inr y=z with totality x z
|
||||
swapMin {x} {y} {z} | inr y=z | inl (inl x<z) = minCommutes x y
|
||||
swapMin {x} {y} {z} | inr y=z | inl (inr z<x) rewrite y=z | minIdempotent z | minCommutes x z = equivMin z<x
|
||||
swapMin {x} {y} {z} | inr y=z | inr x=z = minCommutes x y
|
||||
|
||||
minMin : {x y : A} → (min x (min x y)) ≡ min x y
|
||||
minMin {x} {y} with totality x y
|
||||
minMin {x} {y} | inl (inl x<y) = minIdempotent x
|
||||
minMin {x} {y} | inl (inr y<x) with totality x y
|
||||
minMin {x} {y} | inl (inr y<x) | inl (inl x<y) = exFalso (irreflexive (<Transitive y<x x<y))
|
||||
minMin {x} {y} | inl (inr y<x) | inl (inr y<x') = refl
|
||||
minMin {x} {y} | inl (inr y<x) | inr x=y = x=y
|
||||
minMin {x} {y} | inr x=y = minIdempotent x
|
||||
|
||||
minFromBoth : {l x y : A} → (l < x) → (l < y) → (l < (min x y))
|
||||
minFromBoth {a} {x = x} {y} prX prY with totality x y
|
||||
minFromBoth {a} prX prY | inl (inl x<y) = prX
|
||||
minFromBoth {a} prX prY | inl (inr y<x) = prY
|
||||
minFromBoth {a} prX prY | inr x=y = prX
|
14
Orders/WellFounded/Definition.agda
Normal file
14
Orders/WellFounded/Definition.agda
Normal file
@@ -0,0 +1,14 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Functions
|
||||
|
||||
module Orders.WellFounded.Definition {a b : _} {A : Set a} (_<_ : Rel {a} {b} A) where
|
||||
|
||||
data Accessible (x : A) : Set (lsuc a ⊔ b) where
|
||||
access : (∀ y → y < x → Accessible y) → Accessible x
|
||||
|
||||
WellFounded : Set (lsuc a ⊔ b)
|
||||
WellFounded = ∀ x → Accessible x
|
||||
|
26
Orders/WellFounded/Induction.agda
Normal file
26
Orders/WellFounded/Induction.agda
Normal file
@@ -0,0 +1,26 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Functions
|
||||
open import Orders.WellFounded.Definition
|
||||
|
||||
module Orders.WellFounded.Induction {a b : _} {A : Set a} {_<_ : Rel {a} {b} A} (wf : WellFounded _<_) where
|
||||
|
||||
private
|
||||
foldAcc :
|
||||
{c : _}
|
||||
(P : A → Set c) →
|
||||
(∀ x → (∀ y → y < x → P y) → P x) →
|
||||
∀ z → Accessible _<_ z →
|
||||
P z
|
||||
foldAcc P inductionProof = go
|
||||
where
|
||||
go : (z : A) → (Accessible _<_ z) → P z
|
||||
go z (access prf) = inductionProof z (λ y yLessZ → go y (prf y yLessZ))
|
||||
|
||||
rec :
|
||||
{c : _}
|
||||
(P : A → Set c) →
|
||||
(∀ x → (∀ y → y < x → P y) → P x) → (∀ z → P z)
|
||||
rec P inductionProof z = foldAcc P inductionProof _ (wf z)
|
Reference in New Issue
Block a user