mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-09 13:58:38 +00:00
Hide more stuff (#124)
This commit is contained in:
@@ -26,3 +26,8 @@ boolAnd BoolFalse y = BoolFalse
|
|||||||
boolOr : Bool → Bool → Bool
|
boolOr : Bool → Bool → Bool
|
||||||
boolOr BoolTrue y = BoolTrue
|
boolOr BoolTrue y = BoolTrue
|
||||||
boolOr BoolFalse y = y
|
boolOr BoolFalse y = y
|
||||||
|
|
||||||
|
xor : Bool → Bool → Bool
|
||||||
|
xor BoolTrue BoolTrue = BoolFalse
|
||||||
|
xor BoolTrue BoolFalse = BoolTrue
|
||||||
|
xor BoolFalse b = b
|
||||||
|
17
Boolean/Lemmas.agda
Normal file
17
Boolean/Lemmas.agda
Normal file
@@ -0,0 +1,17 @@
|
|||||||
|
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||||
|
|
||||||
|
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||||
|
open import LogicalFormulae
|
||||||
|
open import Boolean.Definition
|
||||||
|
|
||||||
|
module Boolean.Lemmas where
|
||||||
|
|
||||||
|
notNot : (x : Bool) → not (not x) ≡ x
|
||||||
|
notNot BoolTrue = refl
|
||||||
|
notNot BoolFalse = refl
|
||||||
|
|
||||||
|
notXor : (x y : Bool) → not (xor x y) ≡ xor (not x) y
|
||||||
|
notXor BoolTrue BoolTrue = refl
|
||||||
|
notXor BoolTrue BoolFalse = refl
|
||||||
|
notXor BoolFalse BoolTrue = refl
|
||||||
|
notXor BoolFalse BoolFalse = refl
|
@@ -58,7 +58,6 @@ open import Fields.Orders.Lemmas
|
|||||||
open import Fields.FieldOfFractions.Field
|
open import Fields.FieldOfFractions.Field
|
||||||
open import Fields.FieldOfFractions.Lemmas
|
open import Fields.FieldOfFractions.Lemmas
|
||||||
open import Fields.FieldOfFractions.Order
|
open import Fields.FieldOfFractions.Order
|
||||||
open import Fields.FieldOfFractions.Order
|
|
||||||
--open import Fields.FieldOfFractions.Archimedean
|
--open import Fields.FieldOfFractions.Archimedean
|
||||||
|
|
||||||
open import Rings.Definition
|
open import Rings.Definition
|
||||||
|
@@ -18,24 +18,9 @@ fieldOfFractionsRing : Ring fieldOfFractionsSetoid fieldOfFractionsPlus fieldOfF
|
|||||||
Ring.additiveGroup fieldOfFractionsRing = fieldOfFractionsGroup
|
Ring.additiveGroup fieldOfFractionsRing = fieldOfFractionsGroup
|
||||||
Ring.*WellDefined fieldOfFractionsRing {a} {b} {c} {d} = fieldOfFractionsTimesWellDefined {a} {b} {c} {d}
|
Ring.*WellDefined fieldOfFractionsRing {a} {b} {c} {d} = fieldOfFractionsTimesWellDefined {a} {b} {c} {d}
|
||||||
Ring.1R fieldOfFractionsRing = record { num = Ring.1R R ; denom = Ring.1R R ; denomNonzero = IntegralDomain.nontrivial I }
|
Ring.1R fieldOfFractionsRing = record { num = Ring.1R R ; denom = Ring.1R R ; denomNonzero = IntegralDomain.nontrivial I }
|
||||||
Ring.groupIsAbelian fieldOfFractionsRing {record { num = a ; denom = b }} {record { num = c ; denom = d }} = need
|
Ring.groupIsAbelian fieldOfFractionsRing {record { num = a ; denom = b }} {record { num = c ; denom = d }} = Equivalence.transitive (Setoid.eq S) (Ring.*Commutative R) (Ring.*WellDefined R (Ring.*Commutative R) (Equivalence.transitive (Setoid.eq S) (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) (Ring.*Commutative R)) (Ring.groupIsAbelian R)))
|
||||||
where
|
Ring.*Associative fieldOfFractionsRing {record { num = a ; denom = b }} {record { num = c ; denom = d }} {record { num = e ; denom = f }} = Equivalence.transitive (Setoid.eq S) (Ring.*WellDefined R (Ring.*Associative R) (Ring.*Associative' R)) (Ring.*Commutative R)
|
||||||
open Setoid S
|
Ring.*Commutative fieldOfFractionsRing {record { num = a ; denom = b }} {record { num = c ; denom = d }} = Equivalence.transitive (Setoid.eq S) (Ring.*Commutative R) (Ring.*WellDefined R (Ring.*Commutative R) (Ring.*Commutative R))
|
||||||
open Equivalence eq
|
|
||||||
need : (((a * d) + (b * c)) * (d * b)) ∼ ((b * d) * ((c * b) + (d * a)))
|
|
||||||
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (Ring.*Commutative R) (transitive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) (Ring.*Commutative R)) (Ring.groupIsAbelian R)))
|
|
||||||
Ring.*Associative fieldOfFractionsRing {record { num = a ; denom = b }} {record { num = c ; denom = d }} {record { num = e ; denom = f }} = need
|
|
||||||
where
|
|
||||||
open Setoid S
|
|
||||||
open Equivalence eq
|
|
||||||
need : ((a * (c * e)) * ((b * d) * f)) ∼ ((b * (d * f)) * ((a * c) * e))
|
|
||||||
need = transitive (Ring.*WellDefined R (Ring.*Associative R) (symmetric (Ring.*Associative R))) (Ring.*Commutative R)
|
|
||||||
Ring.*Commutative fieldOfFractionsRing {record { num = a ; denom = b }} {record { num = c ; denom = d }} = need
|
|
||||||
where
|
|
||||||
open Setoid S
|
|
||||||
open Equivalence eq
|
|
||||||
need : ((a * c) * (d * b)) ∼ ((b * d) * (c * a))
|
|
||||||
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (Ring.*Commutative R) (Ring.*Commutative R))
|
|
||||||
Ring.*DistributesOver+ fieldOfFractionsRing {record { num = a ; denom = b }} {record { num = c ; denom = d }} {record { num = e ; denom = f }} = need
|
Ring.*DistributesOver+ fieldOfFractionsRing {record { num = a ; denom = b }} {record { num = c ; denom = d }} {record { num = e ; denom = f }} = need
|
||||||
where
|
where
|
||||||
open Setoid S
|
open Setoid S
|
||||||
|
@@ -12,17 +12,13 @@ open import Semirings.Definition
|
|||||||
open import Functions.Definition
|
open import Functions.Definition
|
||||||
open import Groups.Isomorphisms.Definition
|
open import Groups.Isomorphisms.Definition
|
||||||
open import Boolean.Definition
|
open import Boolean.Definition
|
||||||
|
open import Boolean.Lemmas
|
||||||
|
|
||||||
module Groups.FreeGroup.Parity {a : _} {A : Set a} (decA : DecidableSet A) where
|
module Groups.FreeGroup.Parity {a : _} {A : Set a} (decA : DecidableSet A) where
|
||||||
|
|
||||||
open import Groups.FreeGroup.Word decA
|
open import Groups.FreeGroup.Word decA
|
||||||
open import Groups.FreeGroup.Group decA
|
open import Groups.FreeGroup.Group decA
|
||||||
|
|
||||||
xor : Bool → Bool → Bool
|
|
||||||
xor BoolTrue BoolTrue = BoolFalse
|
|
||||||
xor BoolTrue BoolFalse = BoolTrue
|
|
||||||
xor BoolFalse b = b
|
|
||||||
|
|
||||||
C2 : Group (reflSetoid Bool) xor
|
C2 : Group (reflSetoid Bool) xor
|
||||||
Group.+WellDefined C2 refl refl = refl
|
Group.+WellDefined C2 refl refl = refl
|
||||||
Group.0G C2 = BoolFalse
|
Group.0G C2 = BoolFalse
|
||||||
@@ -39,14 +35,6 @@ Group.invLeft C2 {BoolFalse} = refl
|
|||||||
Group.invRight C2 {BoolTrue} = refl
|
Group.invRight C2 {BoolTrue} = refl
|
||||||
Group.invRight C2 {BoolFalse} = refl
|
Group.invRight C2 {BoolFalse} = refl
|
||||||
|
|
||||||
notNot : (x : Bool) → not (not x) ≡ x
|
|
||||||
notNot BoolTrue = refl
|
|
||||||
notNot BoolFalse = refl
|
|
||||||
|
|
||||||
notWellDefined : {x y : Bool} → (x ≡ y) → not x ≡ not y
|
|
||||||
notWellDefined {BoolTrue} {BoolTrue} x=y = refl
|
|
||||||
notWellDefined {BoolFalse} {BoolFalse} x=y = refl
|
|
||||||
|
|
||||||
parity : (x : A) → ReducedWord → Bool
|
parity : (x : A) → ReducedWord → Bool
|
||||||
parity x empty = BoolFalse
|
parity x empty = BoolFalse
|
||||||
parity x (prependLetter (ofLetter y) w _) with decA x y
|
parity x (prependLetter (ofLetter y) w _) with decA x y
|
||||||
@@ -56,108 +44,103 @@ parity x (prependLetter (ofInv y) w _) with decA x y
|
|||||||
parity x (prependLetter (ofInv y) w _) | inl _ = not (parity x w)
|
parity x (prependLetter (ofInv y) w _) | inl _ = not (parity x w)
|
||||||
parity x (prependLetter (ofInv y) w _) | inr _ = parity x w
|
parity x (prependLetter (ofInv y) w _) | inr _ = parity x w
|
||||||
|
|
||||||
parityPrepend : (a : A) (w : ReducedWord) (l : A) → ((a ≡ l) → False) → parity a (prepend w (ofLetter l)) ≡ parity a w
|
private
|
||||||
parityPrepend a empty l notEq with decA a l
|
parityPrepend : (a : A) (w : ReducedWord) (l : A) → ((a ≡ l) → False) → parity a (prepend w (ofLetter l)) ≡ parity a w
|
||||||
parityPrepend a empty l notEq | inl x = exFalso (notEq x)
|
parityPrepend a empty l notEq with decA a l
|
||||||
parityPrepend a empty l notEq | inr x = refl
|
parityPrepend a empty l notEq | inl x = exFalso (notEq x)
|
||||||
parityPrepend a (prependLetter (ofLetter r) w x) l notEq with decA a l
|
parityPrepend a empty l notEq | inr x = refl
|
||||||
parityPrepend a (prependLetter (ofLetter r) w x) l notEq | inl m = exFalso (notEq m)
|
parityPrepend a (prependLetter (ofLetter r) w x) l notEq with decA a l
|
||||||
parityPrepend a (prependLetter (ofLetter r) w x) l notEq | inr _ = refl
|
parityPrepend a (prependLetter (ofLetter r) w x) l notEq | inl m = exFalso (notEq m)
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq with decA a r
|
parityPrepend a (prependLetter (ofLetter r) w x) l notEq | inr _ = refl
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r with decA l r
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq with decA a r
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inl l=r = exFalso (notEq (transitivity a=r (equalityCommutative l=r)))
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r with decA l r
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r with decA a l
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inl l=r = exFalso (notEq (transitivity a=r (equalityCommutative l=r)))
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r with decA a l
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l with decA a r
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inl x₁ = refl
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l with decA a r
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inr bad = exFalso (bad a=r)
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inl x₁ = refl
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r with decA l r
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inr bad = exFalso (bad a=r)
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl x₁ = refl
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r with decA l r
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r with decA a l
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl x₁ = refl
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r with decA a l
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l with decA a r
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inl a=r = exFalso (a!=r a=r)
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l with decA a r
|
||||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inr x₁ = refl
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inl a=r = exFalso (a!=r a=r)
|
||||||
|
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inr x₁ = refl
|
||||||
|
|
||||||
parityPrepend' : (a : A) (w : ReducedWord) (l : A) → ((a ≡ l) → False) → parity a (prepend w (ofInv l)) ≡ parity a w
|
parityPrepend' : (a : A) (w : ReducedWord) (l : A) → ((a ≡ l) → False) → parity a (prepend w (ofInv l)) ≡ parity a w
|
||||||
parityPrepend' a empty l notEq with decA a l
|
parityPrepend' a empty l notEq with decA a l
|
||||||
parityPrepend' a empty l notEq | inl x = exFalso (notEq x)
|
parityPrepend' a empty l notEq | inl x = exFalso (notEq x)
|
||||||
parityPrepend' a empty l notEq | inr x = refl
|
parityPrepend' a empty l notEq | inr x = refl
|
||||||
parityPrepend' a (prependLetter (ofLetter r) w x) l notEq with decA l r
|
parityPrepend' a (prependLetter (ofLetter r) w x) l notEq with decA l r
|
||||||
parityPrepend' a (prependLetter (ofLetter r) w x) l notEq | inl m with decA a r
|
parityPrepend' a (prependLetter (ofLetter r) w x) l notEq | inl m with decA a r
|
||||||
... | inl a=r = exFalso (notEq (transitivity a=r (equalityCommutative m)))
|
... | inl a=r = exFalso (notEq (transitivity a=r (equalityCommutative m)))
|
||||||
... | inr a!=r = refl
|
... | inr a!=r = refl
|
||||||
parityPrepend' a (prependLetter (ofLetter r) w x) l notEq | inr l!=r with decA a l
|
parityPrepend' a (prependLetter (ofLetter r) w x) l notEq | inr l!=r with decA a l
|
||||||
... | inl a=l = exFalso (notEq a=l)
|
... | inl a=l = exFalso (notEq a=l)
|
||||||
... | inr a!=l = refl
|
... | inr a!=l = refl
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq with decA a r
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq with decA a r
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r with decA l r
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r with decA l r
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inl l=r = exFalso (notEq (transitivity a=r (equalityCommutative l=r)))
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inl l=r = exFalso (notEq (transitivity a=r (equalityCommutative l=r)))
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r with decA a l
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r with decA a l
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l with decA a r
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l with decA a r
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inl x₁ = refl
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inl x₁ = refl
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inr bad = exFalso (bad a=r)
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inr bad = exFalso (bad a=r)
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r with decA l r
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r with decA l r
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl l=r with decA a l
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl l=r with decA a l
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl l=r | inl a=l = exFalso (notEq a=l)
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl l=r | inl a=l = exFalso (notEq a=l)
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl l=r | inr a!=l with decA a r
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl l=r | inr a!=l with decA a r
|
||||||
... | inl a=r = exFalso (a!=r a=r)
|
... | inl a=r = exFalso (a!=r a=r)
|
||||||
... | inr _ = refl
|
... | inr _ = refl
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r with decA a l
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r with decA a l
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l with decA a r
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l with decA a r
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inl a=r = exFalso (a!=r a=r)
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inl a=r = exFalso (a!=r a=r)
|
||||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inr x₁ = refl
|
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inr x₁ = refl
|
||||||
|
|
||||||
parityPrepend'' : (a : A) (w : ReducedWord) → parity a (prepend w (ofLetter a)) ≡ not (parity a w)
|
parityPrepend'' : (a : A) (w : ReducedWord) → parity a (prepend w (ofLetter a)) ≡ not (parity a w)
|
||||||
parityPrepend'' a empty with decA a a
|
parityPrepend'' a empty with decA a a
|
||||||
... | inl _ = refl
|
... | inl _ = refl
|
||||||
... | inr bad = exFalso (bad refl)
|
... | inr bad = exFalso (bad refl)
|
||||||
parityPrepend'' a (prependLetter (ofLetter l) w x) with decA a a
|
parityPrepend'' a (prependLetter (ofLetter l) w x) with decA a a
|
||||||
parityPrepend'' a (prependLetter (ofLetter l) w x) | inl _ with decA a l
|
parityPrepend'' a (prependLetter (ofLetter l) w x) | inl _ with decA a l
|
||||||
parityPrepend'' a (prependLetter (ofLetter l) w x) | inl _ | inl a=l = refl
|
parityPrepend'' a (prependLetter (ofLetter l) w x) | inl _ | inl a=l = refl
|
||||||
parityPrepend'' a (prependLetter (ofLetter l) w x) | inl _ | inr a!=l = refl
|
parityPrepend'' a (prependLetter (ofLetter l) w x) | inl _ | inr a!=l = refl
|
||||||
parityPrepend'' a (prependLetter (ofLetter l) w x) | inr bad = exFalso (bad refl)
|
parityPrepend'' a (prependLetter (ofLetter l) w x) | inr bad = exFalso (bad refl)
|
||||||
parityPrepend'' a (prependLetter (ofInv l) w x) with decA a l
|
parityPrepend'' a (prependLetter (ofInv l) w x) with decA a l
|
||||||
... | inl a=l = equalityCommutative (notNot (parity a w))
|
... | inl a=l = equalityCommutative (notNot (parity a w))
|
||||||
parityPrepend'' a (prependLetter (ofInv l) w x) | inr a!=l with decA a a
|
parityPrepend'' a (prependLetter (ofInv l) w x) | inr a!=l with decA a a
|
||||||
parityPrepend'' a (prependLetter (ofInv l) w x) | inr a!=l | inl _ with decA a l
|
parityPrepend'' a (prependLetter (ofInv l) w x) | inr a!=l | inl _ with decA a l
|
||||||
... | inl a=l = exFalso (a!=l a=l)
|
... | inl a=l = exFalso (a!=l a=l)
|
||||||
... | inr _ = refl
|
... | inr _ = refl
|
||||||
parityPrepend'' a (prependLetter (ofInv l) w x) | inr a!=l | inr bad = exFalso (bad refl)
|
parityPrepend'' a (prependLetter (ofInv l) w x) | inr a!=l | inr bad = exFalso (bad refl)
|
||||||
|
|
||||||
parityPrepend''' : (a : A) (w : ReducedWord) → parity a (prepend w (ofInv a)) ≡ not (parity a w)
|
parityPrepend''' : (a : A) (w : ReducedWord) → parity a (prepend w (ofInv a)) ≡ not (parity a w)
|
||||||
parityPrepend''' a empty with decA a a
|
parityPrepend''' a empty with decA a a
|
||||||
... | inl _ = refl
|
... | inl _ = refl
|
||||||
... | inr bad = exFalso (bad refl)
|
... | inr bad = exFalso (bad refl)
|
||||||
parityPrepend''' a (prependLetter (ofLetter l) w x) with decA a l
|
parityPrepend''' a (prependLetter (ofLetter l) w x) with decA a l
|
||||||
... | inl a=l = equalityCommutative (notNot _)
|
... | inl a=l = equalityCommutative (notNot _)
|
||||||
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l with decA a a
|
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l with decA a a
|
||||||
... | inl _ with decA a l
|
... | inl _ with decA a l
|
||||||
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l | inl _ | inl a=l = exFalso (a!=l a=l)
|
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l | inl _ | inl a=l = exFalso (a!=l a=l)
|
||||||
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l | inl _ | inr _ = refl
|
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l | inl _ | inr _ = refl
|
||||||
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l | inr bad = exFalso (bad refl)
|
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l | inr bad = exFalso (bad refl)
|
||||||
parityPrepend''' a (prependLetter (ofInv l) w x) with decA a a
|
parityPrepend''' a (prependLetter (ofInv l) w x) with decA a a
|
||||||
parityPrepend''' a (prependLetter (ofInv l) w x) | inl _ with decA a l
|
parityPrepend''' a (prependLetter (ofInv l) w x) | inl _ with decA a l
|
||||||
... | inl a=l = refl
|
... | inl a=l = refl
|
||||||
... | inr a!=l = refl
|
... | inr a!=l = refl
|
||||||
parityPrepend''' a (prependLetter (ofInv l) w x) | inr bad = exFalso (bad refl)
|
parityPrepend''' a (prependLetter (ofInv l) w x) | inr bad = exFalso (bad refl)
|
||||||
|
|
||||||
notXor : (x y : Bool) → not (xor x y) ≡ xor (not x) y
|
parityHomIsHom : (a : A) → (x y : ReducedWord) → parity a (_+W_ x y) ≡ xor (parity a x) (parity a y)
|
||||||
notXor BoolTrue BoolTrue = refl
|
parityHomIsHom a empty y = refl
|
||||||
notXor BoolTrue BoolFalse = refl
|
parityHomIsHom a (prependLetter (ofLetter l) x _) y with decA a l
|
||||||
notXor BoolFalse BoolTrue = refl
|
parityHomIsHom a (prependLetter (ofLetter .a) x _) y | inl refl = transitivity (parityPrepend'' a (x +W y)) (transitivity (applyEquality not (parityHomIsHom a x y)) (notXor (parity a x) (parity a y)))
|
||||||
notXor BoolFalse BoolFalse = refl
|
parityHomIsHom a (prependLetter (ofLetter l) x _) y | inr a!=l = transitivity (parityPrepend a (_+W_ x y) l a!=l) (parityHomIsHom a x y)
|
||||||
|
parityHomIsHom a (prependLetter (ofInv l) x _) y with decA a l
|
||||||
parityHomIsHom : (a : A) → (x y : ReducedWord) → parity a (_+W_ x y) ≡ xor (parity a x) (parity a y)
|
parityHomIsHom a (prependLetter (ofInv .a) x _) y | inl refl = transitivity (parityPrepend''' a (x +W y)) (transitivity (applyEquality not (parityHomIsHom a x y)) (notXor (parity a x) (parity a y)))
|
||||||
parityHomIsHom a empty y = refl
|
parityHomIsHom a (prependLetter (ofInv l) x _) y | inr a!=l = transitivity (parityPrepend' a (x +W y) l a!=l) (parityHomIsHom a x y)
|
||||||
parityHomIsHom a (prependLetter (ofLetter l) x _) y with decA a l
|
|
||||||
parityHomIsHom a (prependLetter (ofLetter .a) x _) y | inl refl = transitivity (parityPrepend'' a (x +W y)) (transitivity (notWellDefined (parityHomIsHom a x y)) (notXor (parity a x) (parity a y)))
|
|
||||||
parityHomIsHom a (prependLetter (ofLetter l) x _) y | inr a!=l = transitivity (parityPrepend a (_+W_ x y) l a!=l) (parityHomIsHom a x y)
|
|
||||||
parityHomIsHom a (prependLetter (ofInv l) x _) y with decA a l
|
|
||||||
parityHomIsHom a (prependLetter (ofInv .a) x _) y | inl refl = transitivity (parityPrepend''' a (x +W y)) (transitivity (notWellDefined (parityHomIsHom a x y)) (notXor (parity a x) (parity a y)))
|
|
||||||
parityHomIsHom a (prependLetter (ofInv l) x _) y | inr a!=l = transitivity (parityPrepend' a (x +W y) l a!=l) (parityHomIsHom a x y)
|
|
||||||
|
|
||||||
parityHom : (x : A) → GroupHom (freeGroup) C2 (parity x)
|
parityHom : (x : A) → GroupHom (freeGroup) C2 (parity x)
|
||||||
GroupHom.groupHom (parityHom x) {y} {z} = parityHomIsHom x y z
|
GroupHom.groupHom (parityHom x) {y} {z} = parityHomIsHom x y z
|
||||||
|
@@ -19,60 +19,64 @@ open import Groups.FreeProduct.Setoid decidableIndex decidableGroups G
|
|||||||
open import Groups.FreeProduct.Addition decidableIndex decidableGroups G
|
open import Groups.FreeProduct.Addition decidableIndex decidableGroups G
|
||||||
open import Groups.FreeProduct.Group decidableIndex decidableGroups G
|
open import Groups.FreeProduct.Group decidableIndex decidableGroups G
|
||||||
|
|
||||||
universalPropertyFunction' : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → {i : I} → ReducedSequenceBeginningWith i → C
|
private
|
||||||
universalPropertyFunction' {_+_ = _+_} H fs homs {i} (ofEmpty .i g nonZero) = fs i g
|
universalPropertyFunction' : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → {i : I} → ReducedSequenceBeginningWith i → C
|
||||||
universalPropertyFunction' {_+_ = _+_} H fs homs {i} (prependLetter .i g nonZero x x₁) = (fs i g) + universalPropertyFunction' H fs homs x
|
universalPropertyFunction' {_+_ = _+_} H fs homs {i} (ofEmpty .i g nonZero) = fs i g
|
||||||
|
universalPropertyFunction' {_+_ = _+_} H fs homs {i} (prependLetter .i g nonZero x x₁) = (fs i g) + universalPropertyFunction' H fs homs x
|
||||||
|
|
||||||
universalPropertyFunction : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → ReducedSequence → C
|
universalPropertyFunction : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → ReducedSequence → C
|
||||||
universalPropertyFunction H fs homs empty = Group.0G H
|
universalPropertyFunction H fs homs empty = Group.0G H
|
||||||
universalPropertyFunction H fs homs (nonempty i x) = universalPropertyFunction' H fs homs x
|
universalPropertyFunction H fs homs (nonempty i x) = universalPropertyFunction' H fs homs x
|
||||||
|
|
||||||
upWellDefined' : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → {m n : I} → (x : ReducedSequenceBeginningWith m) (y : ReducedSequenceBeginningWith n) → (eq : =RP' x y) → Setoid._∼_ T (universalPropertyFunction H fs homs (nonempty m x)) (universalPropertyFunction H fs homs (nonempty n y))
|
private
|
||||||
upWellDefined' H fs homs (ofEmpty m g nonZero) (ofEmpty n g₁ nonZero₁) eq with decidableIndex m n
|
upWellDefined' : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → {m n : I} → (x : ReducedSequenceBeginningWith m) (y : ReducedSequenceBeginningWith n) → (eq : =RP' x y) → Setoid._∼_ T (universalPropertyFunction H fs homs (nonempty m x)) (universalPropertyFunction H fs homs (nonempty n y))
|
||||||
... | inl refl = GroupHom.wellDefined (homs m) eq
|
upWellDefined' H fs homs (ofEmpty m g nonZero) (ofEmpty n g₁ nonZero₁) eq with decidableIndex m n
|
||||||
upWellDefined' H fs homs (prependLetter m g nonZero x x₁) (prependLetter n g₁ nonZero₁ y x₂) eq with decidableIndex m n
|
... | inl refl = GroupHom.wellDefined (homs m) eq
|
||||||
... | inl refl = Group.+WellDefined H (GroupHom.wellDefined (homs m) (_&&_.fst eq)) (upWellDefined' H fs homs x y (_&&_.snd eq))
|
upWellDefined' H fs homs (prependLetter m g nonZero x x₁) (prependLetter n g₁ nonZero₁ y x₂) eq with decidableIndex m n
|
||||||
|
... | inl refl = Group.+WellDefined H (GroupHom.wellDefined (homs m) (_&&_.fst eq)) (upWellDefined' H fs homs x y (_&&_.snd eq))
|
||||||
|
|
||||||
upWellDefined : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → (x : ReducedSequence) (y : ReducedSequence) → (eq : _=RP_ x y) → Setoid._∼_ T (universalPropertyFunction H fs homs x) (universalPropertyFunction H fs homs y)
|
upWellDefined : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → (x : ReducedSequence) (y : ReducedSequence) → (eq : _=RP_ x y) → Setoid._∼_ T (universalPropertyFunction H fs homs x) (universalPropertyFunction H fs homs y)
|
||||||
upWellDefined {T = T} H fs homs empty empty eq = Equivalence.reflexive (Setoid.eq T)
|
upWellDefined {T = T} H fs homs empty empty eq = Equivalence.reflexive (Setoid.eq T)
|
||||||
upWellDefined H fs homs (nonempty i w1) (nonempty j w2) eq = upWellDefined' H fs homs w1 w2 eq
|
upWellDefined H fs homs (nonempty i w1) (nonempty j w2) eq = upWellDefined' H fs homs w1 w2 eq
|
||||||
|
|
||||||
upPrepend : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → {j : I} (y : ReducedSequence) → (g : A j) .(pr : _) → Setoid._∼_ T (universalPropertyFunction H fs homs (prepend j g pr y)) ((fs j g) + universalPropertyFunction H fs homs y)
|
private
|
||||||
upPrepend {T = T} H fs homs empty g pr = Equivalence.symmetric (Setoid.eq T) (Group.identRight H)
|
upPrepend : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → {j : I} (y : ReducedSequence) → (g : A j) .(pr : _) → Setoid._∼_ T (universalPropertyFunction H fs homs (prepend j g pr y)) ((fs j g) + universalPropertyFunction H fs homs y)
|
||||||
upPrepend {T = T} H fs homs {j} (nonempty i (ofEmpty .i h nonZero)) g pr with decidableIndex j i
|
upPrepend {T = T} H fs homs empty g pr = Equivalence.symmetric (Setoid.eq T) (Group.identRight H)
|
||||||
... | inr j!=i = Equivalence.reflexive (Setoid.eq T)
|
upPrepend {T = T} H fs homs {j} (nonempty i (ofEmpty .i h nonZero)) g pr with decidableIndex j i
|
||||||
... | inl refl with decidableGroups j ((j + g) h) (Group.0G (G j))
|
... | inr j!=i = Equivalence.reflexive (Setoid.eq T)
|
||||||
... | inl x = Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (imageOfIdentityIsIdentity (homs j))) (Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (GroupHom.wellDefined (homs j) x)) (GroupHom.groupHom (homs j)))
|
... | inl refl with decidableGroups j ((j + g) h) (Group.0G (G j))
|
||||||
... | inr x = GroupHom.groupHom (homs j)
|
... | inl x = Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (imageOfIdentityIsIdentity (homs j))) (Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (GroupHom.wellDefined (homs j) x)) (GroupHom.groupHom (homs j)))
|
||||||
upPrepend {T = T} H fs homs {j} (nonempty k (prependLetter .k h nonZero y _)) g pr with decidableIndex j k
|
... | inr x = GroupHom.groupHom (homs j)
|
||||||
... | inr j!=k = Equivalence.reflexive (Setoid.eq T)
|
upPrepend {T = T} H fs homs {j} (nonempty k (prependLetter .k h nonZero y _)) g pr with decidableIndex j k
|
||||||
... | inl refl with decidableGroups j ((j + g) h) (Group.0G (G j))
|
... | inr j!=k = Equivalence.reflexive (Setoid.eq T)
|
||||||
... | inl x = transitive (symmetric (Group.identLeft H)) (transitive (Group.+WellDefined H (transitive (symmetric (imageOfIdentityIsIdentity (homs k))) (transitive (GroupHom.wellDefined (homs k) (Equivalence.symmetric (Setoid.eq (S k)) x)) (GroupHom.groupHom (homs k)))) reflexive) (symmetric (Group.+Associative H)))
|
... | inl refl with decidableGroups j ((j + g) h) (Group.0G (G j))
|
||||||
where
|
... | inl x = transitive (symmetric (Group.identLeft H)) (transitive (Group.+WellDefined H (transitive (symmetric (imageOfIdentityIsIdentity (homs k))) (transitive (GroupHom.wellDefined (homs k) (Equivalence.symmetric (Setoid.eq (S k)) x)) (GroupHom.groupHom (homs k)))) reflexive) (symmetric (Group.+Associative H)))
|
||||||
open Setoid T
|
where
|
||||||
open Equivalence eq
|
open Setoid T
|
||||||
... | inr x = transitive (Group.+WellDefined H (GroupHom.groupHom (homs k)) reflexive) (symmetric (Group.+Associative H))
|
open Equivalence eq
|
||||||
where
|
... | inr x = transitive (Group.+WellDefined H (GroupHom.groupHom (homs k)) reflexive) (symmetric (Group.+Associative H))
|
||||||
open Setoid T
|
where
|
||||||
open Equivalence eq
|
open Setoid T
|
||||||
|
open Equivalence eq
|
||||||
|
|
||||||
upHom : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → {i : I} (x : ReducedSequenceBeginningWith i) (y : ReducedSequence) → Setoid._∼_ T (universalPropertyFunction H fs homs (plus' x y)) (universalPropertyFunction' H fs homs x + universalPropertyFunction H fs homs y)
|
private
|
||||||
upHom {T = T} H fs homs (ofEmpty _ g nonZero) empty = Equivalence.symmetric (Setoid.eq T) (Group.identRight H)
|
upHom : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → {i : I} (x : ReducedSequenceBeginningWith i) (y : ReducedSequence) → Setoid._∼_ T (universalPropertyFunction H fs homs (plus' x y)) (universalPropertyFunction' H fs homs x + universalPropertyFunction H fs homs y)
|
||||||
upHom {T = T} H fs homs (ofEmpty j g nonZero) (nonempty i (ofEmpty .i h nonZero1)) with decidableIndex j i
|
upHom {T = T} H fs homs (ofEmpty _ g nonZero) empty = Equivalence.symmetric (Setoid.eq T) (Group.identRight H)
|
||||||
... | inr j!=i = Equivalence.reflexive (Setoid.eq T)
|
upHom {T = T} H fs homs (ofEmpty j g nonZero) (nonempty i (ofEmpty .i h nonZero1)) with decidableIndex j i
|
||||||
... | inl refl with decidableGroups j ((j + g) h) (Group.0G (G j))
|
... | inr j!=i = Equivalence.reflexive (Setoid.eq T)
|
||||||
... | inl x = Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (imageOfIdentityIsIdentity (homs j))) (Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (GroupHom.wellDefined (homs j) x)) (GroupHom.groupHom (homs j)))
|
... | inl refl with decidableGroups j ((j + g) h) (Group.0G (G j))
|
||||||
... | inr x = GroupHom.groupHom (homs j)
|
... | inl x = Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (imageOfIdentityIsIdentity (homs j))) (Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (GroupHom.wellDefined (homs j) x)) (GroupHom.groupHom (homs j)))
|
||||||
upHom {T = T} H fs homs (ofEmpty j g nonZero) (nonempty i (prependLetter .i h nonZero1 x x₁)) with decidableIndex j i
|
... | inr x = GroupHom.groupHom (homs j)
|
||||||
... | inr j!=i = Equivalence.reflexive (Setoid.eq T)
|
upHom {T = T} H fs homs (ofEmpty j g nonZero) (nonempty i (prependLetter .i h nonZero1 x x₁)) with decidableIndex j i
|
||||||
... | inl refl with decidableGroups j ((j + g) h) (Group.0G (G j))
|
... | inr j!=i = Equivalence.reflexive (Setoid.eq T)
|
||||||
... | inr _ = Equivalence.transitive (Setoid.eq T) (Group.+WellDefined H (GroupHom.groupHom (homs j)) (Equivalence.reflexive (Setoid.eq T))) (Equivalence.symmetric (Setoid.eq T) (Group.+Associative H))
|
... | inl refl with decidableGroups j ((j + g) h) (Group.0G (G j))
|
||||||
... | inl eq1 = Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (Group.identLeft H)) (Equivalence.transitive (Setoid.eq T) (Group.+WellDefined H (Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (imageOfIdentityIsIdentity (homs j))) (Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined (homs j) (Equivalence.symmetric (Setoid.eq (S j)) eq1)) (GroupHom.groupHom (homs j)))) (Equivalence.reflexive (Setoid.eq T))) (Equivalence.symmetric (Setoid.eq T) (Group.+Associative H)))
|
... | inr _ = Equivalence.transitive (Setoid.eq T) (Group.+WellDefined H (GroupHom.groupHom (homs j)) (Equivalence.reflexive (Setoid.eq T))) (Equivalence.symmetric (Setoid.eq T) (Group.+Associative H))
|
||||||
upHom {T = T} H fs homs (prependLetter j g nonZero {k} w k!=j) empty = Equivalence.transitive (Setoid.eq T) (Equivalence.transitive (Setoid.eq T) (upWellDefined H fs homs (plus' (prependLetter j g _ w k!=j) empty) (prepend j g _ (nonempty k w)) (prependWD' g nonZero (plus' w empty) (nonempty k w) (plusEmptyRight w))) (upPrepend H fs homs (nonempty k w) g nonZero)) (Equivalence.symmetric (Setoid.eq T) (Group.identRight H))
|
... | inl eq1 = Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (Group.identLeft H)) (Equivalence.transitive (Setoid.eq T) (Group.+WellDefined H (Equivalence.transitive (Setoid.eq T) (Equivalence.symmetric (Setoid.eq T) (imageOfIdentityIsIdentity (homs j))) (Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined (homs j) (Equivalence.symmetric (Setoid.eq (S j)) eq1)) (GroupHom.groupHom (homs j)))) (Equivalence.reflexive (Setoid.eq T))) (Equivalence.symmetric (Setoid.eq T) (Group.+Associative H)))
|
||||||
upHom {T = T} H fs homs (prependLetter j g nonZero {k} m k!=j) (nonempty i x2) = transitive (upPrepend H fs homs (plus' m (nonempty i x2)) g nonZero) (transitive (Group.+WellDefined H reflexive (upHom H fs homs m (nonempty i x2))) (Group.+Associative H))
|
upHom {T = T} H fs homs (prependLetter j g nonZero {k} w k!=j) empty = Equivalence.transitive (Setoid.eq T) (Equivalence.transitive (Setoid.eq T) (upWellDefined H fs homs (plus' (prependLetter j g _ w k!=j) empty) (prepend j g _ (nonempty k w)) (prependWD' g nonZero (plus' w empty) (nonempty k w) (plusEmptyRight w))) (upPrepend H fs homs (nonempty k w) g nonZero)) (Equivalence.symmetric (Setoid.eq T) (Group.identRight H))
|
||||||
where
|
upHom {T = T} H fs homs (prependLetter j g nonZero {k} m k!=j) (nonempty i x2) = transitive (upPrepend H fs homs (plus' m (nonempty i x2)) g nonZero) (transitive (Group.+WellDefined H reflexive (upHom H fs homs m (nonempty i x2))) (Group.+Associative H))
|
||||||
open Setoid T
|
where
|
||||||
open Equivalence eq
|
open Setoid T
|
||||||
|
open Equivalence eq
|
||||||
|
|
||||||
universalPropertyHom : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → GroupHom FreeProductGroup H (universalPropertyFunction H fs homs)
|
universalPropertyHom : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → GroupHom FreeProductGroup H (universalPropertyFunction H fs homs)
|
||||||
GroupHom.wellDefined (universalPropertyHom {T = T} H fs homs) {x} {y} eq = upWellDefined H fs homs x y eq
|
GroupHom.wellDefined (universalPropertyHom {T = T} H fs homs) {x} {y} eq = upWellDefined H fs homs x y eq
|
||||||
@@ -86,29 +90,30 @@ GroupHom.groupHom (universalPropertyHom H fs homs) {nonempty i x} {nonempty j y}
|
|||||||
universalPropertyFunctionHasProperty : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → {i : I} → (g : A i) → (nz : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → Setoid._∼_ T (fs i g) (universalPropertyFunction H fs homs (injection g nz))
|
universalPropertyFunctionHasProperty : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → {i : I} → (g : A i) → (nz : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → Setoid._∼_ T (fs i g) (universalPropertyFunction H fs homs (injection g nz))
|
||||||
universalPropertyFunctionHasProperty {T = T} H fs homs g nz = Equivalence.reflexive (Setoid.eq T)
|
universalPropertyFunctionHasProperty {T = T} H fs homs g nz = Equivalence.reflexive (Setoid.eq T)
|
||||||
|
|
||||||
universalPropertyFunctionUniquelyHasPropertyLemma : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → (otherFunction : ReducedSequence → C) → (isHom : GroupHom FreeProductGroup H otherFunction) → ({i : I} → (g : A i) → .(nz : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → Setoid._∼_ T (fs i g) (otherFunction (injection g nz))) → {k l : I} (neq : (k ≡ l) → False) (r : ReducedSequenceBeginningWith l) (g : A k) .(nz : (Setoid._∼_ (S k) g (Group.0G (G k)) → False)) → Setoid._∼_ T (otherFunction (nonempty k (prependLetter k g nz r neq))) (fs k g + universalPropertyFunction' H fs homs r)
|
private
|
||||||
universalPropertyFunctionUniquelyHasPropertyLemma {T = T} H fs homs otherFunction hom x {k} {l} neq (ofEmpty .l g2 nonZero) g nz = transitive (GroupHom.wellDefined hom {nonempty k (prependLetter k g nz (ofEmpty l g2 nonZero) neq)} {nonempty _ (ofEmpty k g nz) +RP nonempty _ (ofEmpty l g2 nonZero)} t) (transitive (GroupHom.groupHom hom {nonempty k (ofEmpty k g nz)} {nonempty _ (ofEmpty l g2 nonZero)}) (Group.+WellDefined H (symmetric (x g nz)) (symmetric (x g2 nonZero))))
|
universalPropertyFunctionUniquelyHasPropertyLemma : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → (otherFunction : ReducedSequence → C) → (isHom : GroupHom FreeProductGroup H otherFunction) → ({i : I} → (g : A i) → .(nz : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → Setoid._∼_ T (fs i g) (otherFunction (injection g nz))) → {k l : I} (neq : (k ≡ l) → False) (r : ReducedSequenceBeginningWith l) (g : A k) .(nz : (Setoid._∼_ (S k) g (Group.0G (G k)) → False)) → Setoid._∼_ T (otherFunction (nonempty k (prependLetter k g nz r neq))) (fs k g + universalPropertyFunction' H fs homs r)
|
||||||
where
|
universalPropertyFunctionUniquelyHasPropertyLemma {T = T} H fs homs otherFunction hom x {k} {l} neq (ofEmpty .l g2 nonZero) g nz = transitive (GroupHom.wellDefined hom {nonempty k (prependLetter k g nz (ofEmpty l g2 nonZero) neq)} {nonempty _ (ofEmpty k g nz) +RP nonempty _ (ofEmpty l g2 nonZero)} t) (transitive (GroupHom.groupHom hom {nonempty k (ofEmpty k g nz)} {nonempty _ (ofEmpty l g2 nonZero)}) (Group.+WellDefined H (symmetric (x g nz)) (symmetric (x g2 nonZero))))
|
||||||
open Setoid T
|
where
|
||||||
open Equivalence eq
|
open Setoid T
|
||||||
t : Setoid._∼_ freeProductSetoid (nonempty k (prependLetter k g nz (ofEmpty l g2 nonZero) neq)) (prepend k g nz (nonempty l (ofEmpty l g2 nonZero)))
|
open Equivalence eq
|
||||||
t with decidableIndex k l
|
t : Setoid._∼_ freeProductSetoid (nonempty k (prependLetter k g nz (ofEmpty l g2 nonZero) neq)) (prepend k g nz (nonempty l (ofEmpty l g2 nonZero)))
|
||||||
... | inl p = exFalso (neq p)
|
t with decidableIndex k l
|
||||||
... | inr _ with decidableIndex k k
|
... | inl p = exFalso (neq p)
|
||||||
... | inr bad = exFalso (bad refl)
|
... | inr _ with decidableIndex k k
|
||||||
... | inl refl = Equivalence.reflexive (Setoid.eq (S k)) ,, =RP'reflex (ofEmpty l g2 _)
|
... | inr bad = exFalso (bad refl)
|
||||||
universalPropertyFunctionUniquelyHasPropertyLemma {T = T} H fs homs otherFunction hom x {k} {l} neq (prependLetter .l h nonZero r pr) g nz = transitive (GroupHom.wellDefined hom {nonempty _ (prependLetter k g nz (prependLetter l h nonZero r pr) neq)} {(nonempty k (ofEmpty k g nz)) +RP (nonempty l (prependLetter l h nonZero r pr))} t) (transitive (GroupHom.groupHom hom {nonempty k (ofEmpty k g nz)} {nonempty l (prependLetter l h nonZero r pr)}) (Group.+WellDefined H (symmetric (x g nz)) (universalPropertyFunctionUniquelyHasPropertyLemma H fs homs otherFunction hom x pr r h nonZero)))
|
... | inl refl = Equivalence.reflexive (Setoid.eq (S k)) ,, =RP'reflex (ofEmpty l g2 _)
|
||||||
where
|
universalPropertyFunctionUniquelyHasPropertyLemma {T = T} H fs homs otherFunction hom x {k} {l} neq (prependLetter .l h nonZero r pr) g nz = transitive (GroupHom.wellDefined hom {nonempty _ (prependLetter k g nz (prependLetter l h nonZero r pr) neq)} {(nonempty k (ofEmpty k g nz)) +RP (nonempty l (prependLetter l h nonZero r pr))} t) (transitive (GroupHom.groupHom hom {nonempty k (ofEmpty k g nz)} {nonempty l (prependLetter l h nonZero r pr)}) (Group.+WellDefined H (symmetric (x g nz)) (universalPropertyFunctionUniquelyHasPropertyLemma H fs homs otherFunction hom x pr r h nonZero)))
|
||||||
open Setoid T
|
where
|
||||||
open Equivalence eq
|
open Setoid T
|
||||||
t : Setoid._∼_ freeProductSetoid (nonempty k (prependLetter k g nz (prependLetter l h nonZero r pr) neq)) (prepend k g nz (nonempty l (prependLetter l h nonZero r pr)))
|
open Equivalence eq
|
||||||
t with decidableIndex k l
|
t : Setoid._∼_ freeProductSetoid (nonempty k (prependLetter k g nz (prependLetter l h nonZero r pr) neq)) (prepend k g nz (nonempty l (prependLetter l h nonZero r pr)))
|
||||||
... | inl bad = exFalso (neq bad)
|
t with decidableIndex k l
|
||||||
... | inr k!=l with decidableIndex k k
|
... | inl bad = exFalso (neq bad)
|
||||||
... | inr bad = exFalso (bad refl)
|
... | inr k!=l with decidableIndex k k
|
||||||
... | inl refl with decidableIndex l l
|
... | inr bad = exFalso (bad refl)
|
||||||
... | inr bad = exFalso (bad refl)
|
... | inl refl with decidableIndex l l
|
||||||
... | inl refl = Equivalence.reflexive (Setoid.eq (S k)) ,, ((Equivalence.reflexive (Setoid.eq (S l))) ,, =RP'reflex r)
|
... | inr bad = exFalso (bad refl)
|
||||||
|
... | inl refl = Equivalence.reflexive (Setoid.eq (S k)) ,, ((Equivalence.reflexive (Setoid.eq (S l))) ,, =RP'reflex r)
|
||||||
|
|
||||||
universalPropertyFunctionUniquelyHasProperty : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → (otherFunction : ReducedSequence → C) → (isHom : GroupHom FreeProductGroup H otherFunction) → ({i : I} → (g : A i) → .(nz : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → Setoid._∼_ T (fs i g) (otherFunction (injection g nz))) → (r : ReducedSequence) → Setoid._∼_ T (otherFunction r) (universalPropertyFunction H fs homs r)
|
universalPropertyFunctionUniquelyHasProperty : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C → C} (H : Group T _+_) → (fs : (i : I) → (A i → C)) → (homs : (i : I) → GroupHom (G i) H (fs i)) → (otherFunction : ReducedSequence → C) → (isHom : GroupHom FreeProductGroup H otherFunction) → ({i : I} → (g : A i) → .(nz : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → Setoid._∼_ T (fs i g) (otherFunction (injection g nz))) → (r : ReducedSequence) → Setoid._∼_ T (otherFunction r) (universalPropertyFunction H fs homs r)
|
||||||
universalPropertyFunctionUniquelyHasProperty H fs homs otherFunction hom prop empty = imageOfIdentityIsIdentity hom
|
universalPropertyFunctionUniquelyHasProperty H fs homs otherFunction hom prop empty = imageOfIdentityIsIdentity hom
|
||||||
|
@@ -18,14 +18,15 @@ _*B_ : BinNat → BinNat → BinNat
|
|||||||
(zero :: a) *B b = zero :: (a *B b)
|
(zero :: a) *B b = zero :: (a *B b)
|
||||||
(one :: a) *B b = (zero :: (a *B b)) +B b
|
(one :: a) *B b = (zero :: (a *B b)) +B b
|
||||||
|
|
||||||
contr : {a : _} {A : Set a} {l1 l2 : List A} → {x : A} → l1 ≡ [] → l1 ≡ x :: l2 → False
|
private
|
||||||
contr {l1 = []} p1 ()
|
contr : {a : _} {A : Set a} {l1 l2 : List A} → {x : A} → l1 ≡ [] → l1 ≡ x :: l2 → False
|
||||||
contr {l1 = x :: l1} () p2
|
contr {l1 = []} p1 ()
|
||||||
|
contr {l1 = x :: l1} () p2
|
||||||
|
|
||||||
*BEmpty : (a : BinNat) → canonical (a *B []) ≡ []
|
*BEmpty : (a : BinNat) → canonical (a *B []) ≡ []
|
||||||
*BEmpty [] = refl
|
*BEmpty [] = refl
|
||||||
*BEmpty (zero :: a) rewrite *BEmpty a = refl
|
*BEmpty (zero :: a) rewrite *BEmpty a = refl
|
||||||
*BEmpty (one :: a) rewrite *BEmpty a = refl
|
*BEmpty (one :: a) rewrite *BEmpty a = refl
|
||||||
|
|
||||||
canonicalDistributesPlus : (a b : BinNat) → canonical (a +B b) ≡ canonical a +B canonical b
|
canonicalDistributesPlus : (a b : BinNat) → canonical (a +B b) ≡ canonical a +B canonical b
|
||||||
canonicalDistributesPlus a b = transitivity ans (+BIsInherited (canonical a) (canonical b) (canonicalIdempotent a) (canonicalIdempotent b))
|
canonicalDistributesPlus a b = transitivity ans (+BIsInherited (canonical a) (canonical b) (canonicalIdempotent a) (canonicalIdempotent b))
|
||||||
|
@@ -11,11 +11,12 @@ open import Semirings.Definition
|
|||||||
|
|
||||||
module Numbers.BinaryNaturals.Order where
|
module Numbers.BinaryNaturals.Order where
|
||||||
|
|
||||||
data Compare : Set where
|
data Compare : Set where
|
||||||
Equal : Compare
|
Equal : Compare
|
||||||
FirstLess : Compare
|
FirstLess : Compare
|
||||||
FirstGreater : Compare
|
FirstGreater : Compare
|
||||||
|
|
||||||
|
private
|
||||||
badCompare : Equal ≡ FirstLess → False
|
badCompare : Equal ≡ FirstLess → False
|
||||||
badCompare ()
|
badCompare ()
|
||||||
|
|
||||||
@@ -25,12 +26,13 @@ module Numbers.BinaryNaturals.Order where
|
|||||||
badCompare'' : FirstLess ≡ FirstGreater → False
|
badCompare'' : FirstLess ≡ FirstGreater → False
|
||||||
badCompare'' ()
|
badCompare'' ()
|
||||||
|
|
||||||
_<BInherited_ : BinNat → BinNat → Compare
|
_<BInherited_ : BinNat → BinNat → Compare
|
||||||
a <BInherited b with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
a <BInherited b with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||||
(a <BInherited b) | inl (inl x) = FirstLess
|
(a <BInherited b) | inl (inl x) = FirstLess
|
||||||
(a <BInherited b) | inl (inr x) = FirstGreater
|
(a <BInherited b) | inl (inr x) = FirstGreater
|
||||||
(a <BInherited b) | inr x = Equal
|
(a <BInherited b) | inr x = Equal
|
||||||
|
|
||||||
|
private
|
||||||
go<B : Compare → BinNat → BinNat → Compare
|
go<B : Compare → BinNat → BinNat → Compare
|
||||||
go<B Equal [] [] = Equal
|
go<B Equal [] [] = Equal
|
||||||
go<B Equal [] (zero :: b) = go<B Equal [] b
|
go<B Equal [] (zero :: b) = go<B Equal [] b
|
||||||
@@ -58,9 +60,10 @@ module Numbers.BinaryNaturals.Order where
|
|||||||
go<B FirstLess (one :: a) (zero :: b) = go<B FirstGreater a b
|
go<B FirstLess (one :: a) (zero :: b) = go<B FirstGreater a b
|
||||||
go<B FirstLess (one :: a) (one :: b) = go<B FirstLess a b
|
go<B FirstLess (one :: a) (one :: b) = go<B FirstLess a b
|
||||||
|
|
||||||
_<B_ : BinNat → BinNat → Compare
|
_<B_ : BinNat → BinNat → Compare
|
||||||
a <B b = go<B Equal a b
|
a <B b = go<B Equal a b
|
||||||
|
|
||||||
|
private
|
||||||
lemma1 : {s : Compare} → (n : BinNat) → go<B s n n ≡ s
|
lemma1 : {s : Compare} → (n : BinNat) → go<B s n n ≡ s
|
||||||
lemma1 {Equal} [] = refl
|
lemma1 {Equal} [] = refl
|
||||||
lemma1 {Equal} (zero :: n) = lemma1 n
|
lemma1 {Equal} (zero :: n) = lemma1 n
|
||||||
@@ -304,73 +307,73 @@ module Numbers.BinaryNaturals.Order where
|
|||||||
chopDouble a b one | inr a=b | inl (inl a<b) rewrite a=b = exFalso (TotalOrder.irreflexive (ℕTotalOrder) a<b)
|
chopDouble a b one | inr a=b | inl (inl a<b) rewrite a=b = exFalso (TotalOrder.irreflexive (ℕTotalOrder) a<b)
|
||||||
chopDouble a b zero | inr a=b | inl (inr b<a) rewrite a=b = exFalso (TotalOrder.irreflexive (ℕTotalOrder) b<a)
|
chopDouble a b zero | inr a=b | inl (inr b<a) rewrite a=b = exFalso (TotalOrder.irreflexive (ℕTotalOrder) b<a)
|
||||||
chopDouble a b one | inr a=b | inl (inr b<a) rewrite a=b = exFalso (TotalOrder.irreflexive (ℕTotalOrder) b<a)
|
chopDouble a b one | inr a=b | inl (inr b<a) rewrite a=b = exFalso (TotalOrder.irreflexive (ℕTotalOrder) b<a)
|
||||||
chopDouble a b i | inr a=b | inr x = refl
|
chopDouble a b _ | inr a=b | inr x = refl
|
||||||
|
|
||||||
succNotLess : {n : ℕ} → succ n <N n → False
|
succNotLess : {n : ℕ} → succ n <N n → False
|
||||||
succNotLess {succ n} (le x proof) = succNotLess {n} (le x (succInjective (transitivity (applyEquality succ (transitivity (Semiring.commutative ℕSemiring (succ x) (succ n)) (transitivity (applyEquality succ (transitivity (Semiring.commutative ℕSemiring n (succ x)) (applyEquality succ (Semiring.commutative ℕSemiring x n)))) (Semiring.commutative ℕSemiring (succ (succ n)) x)))) proof)))
|
succNotLess {succ n} (le x proof) = succNotLess {n} (le x (succInjective (transitivity (applyEquality succ (transitivity (Semiring.commutative ℕSemiring (succ x) (succ n)) (transitivity (applyEquality succ (transitivity (Semiring.commutative ℕSemiring n (succ x)) (applyEquality succ (Semiring.commutative ℕSemiring x n)))) (Semiring.commutative ℕSemiring (succ (succ n)) x)))) proof)))
|
||||||
|
|
||||||
<BIsInherited : (a b : BinNat) → a <BInherited b ≡ a <B b
|
<BIsInherited : (a b : BinNat) → a <BInherited b ≡ a <B b
|
||||||
<BIsInherited [] b with TotalOrder.totality ℕTotalOrder 0 (binNatToN b)
|
<BIsInherited [] b with TotalOrder.totality ℕTotalOrder 0 (binNatToN b)
|
||||||
<BIsInherited [] b | inl (inl x) with inspect (binNatToN b)
|
<BIsInherited [] b | inl (inl x) with inspect (binNatToN b)
|
||||||
<BIsInherited [] b | inl (inl x) | 0 with≡ pr rewrite binNatToNZero b pr | pr = exFalso (TotalOrder.irreflexive (ℕTotalOrder) x)
|
<BIsInherited [] b | inl (inl x) | 0 with≡ pr rewrite binNatToNZero b pr | pr = exFalso (TotalOrder.irreflexive (ℕTotalOrder) x)
|
||||||
<BIsInherited [] b | inl (inl x) | (succ bl) with≡ pr rewrite pr = equalityCommutative (zeroLess b λ p → zeroNotSucc bl b p pr)
|
<BIsInherited [] b | inl (inl x) | (succ bl) with≡ pr rewrite pr = equalityCommutative (zeroLess b λ p → zeroNotSucc bl b p pr)
|
||||||
<BIsInherited [] b | inr 0=b rewrite canonicalSecond [] b Equal | binNatToNZero b (equalityCommutative 0=b) = refl
|
<BIsInherited [] b | inr 0=b rewrite canonicalSecond [] b Equal | binNatToNZero b (equalityCommutative 0=b) = refl
|
||||||
<BIsInherited (a :: as) [] with TotalOrder.totality ℕTotalOrder (binNatToN (a :: as)) 0
|
<BIsInherited (a :: as) [] with TotalOrder.totality ℕTotalOrder (binNatToN (a :: as)) 0
|
||||||
<BIsInherited (a :: as) [] | inl (inr x) with inspect (binNatToN (a :: as))
|
<BIsInherited (a :: as) [] | inl (inr x) with inspect (binNatToN (a :: as))
|
||||||
<BIsInherited (a :: as) [] | inl (inr x) | zero with≡ pr rewrite binNatToNZero (a :: as) pr | pr = exFalso (TotalOrder.irreflexive (ℕTotalOrder) x)
|
<BIsInherited (a :: as) [] | inl (inr x) | zero with≡ pr rewrite binNatToNZero (a :: as) pr | pr = exFalso (TotalOrder.irreflexive (ℕTotalOrder) x)
|
||||||
<BIsInherited (a :: as) [] | inl (inr x) | succ y with≡ pr rewrite pr = equalityCommutative (zeroLess' (a :: as) λ i → zeroNotSucc y (a :: as) i pr)
|
<BIsInherited (a :: as) [] | inl (inr x) | succ y with≡ pr rewrite pr = equalityCommutative (zeroLess' (a :: as) λ i → zeroNotSucc y (a :: as) i pr)
|
||||||
<BIsInherited (a :: as) [] | inr x rewrite canonicalFirst (a :: as) [] Equal | binNatToNZero (a :: as) x = refl
|
<BIsInherited (a :: as) [] | inr x rewrite canonicalFirst (a :: as) [] Equal | binNatToNZero (a :: as) x = refl
|
||||||
<BIsInherited (zero :: a) (zero :: b) = transitivity (chopDouble a b zero) (<BIsInherited a b)
|
<BIsInherited (zero :: a) (zero :: b) = transitivity (chopDouble a b zero) (<BIsInherited a b)
|
||||||
<BIsInherited (zero :: a) (one :: b) with TotalOrder.totality ℕTotalOrder (binNatToN (zero :: a)) (binNatToN (one :: b))
|
<BIsInherited (zero :: a) (one :: b) with TotalOrder.totality ℕTotalOrder (binNatToN (zero :: a)) (binNatToN (one :: b))
|
||||||
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||||
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) | inl (inl a<b) = equalityCommutative (equalToFirstLess FirstLess a b (equalityCommutative indHyp))
|
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) | inl (inl a<b) = equalityCommutative (equalToFirstLess FirstLess a b (equalityCommutative indHyp))
|
||||||
where
|
|
||||||
t : a <BInherited b ≡ FirstLess
|
|
||||||
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
|
||||||
t | inl (inl x) = refl
|
|
||||||
t | inl (inr x) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) x a<b))
|
|
||||||
t | inr x rewrite x = exFalso (TotalOrder.irreflexive (ℕTotalOrder) a<b)
|
|
||||||
indHyp : FirstLess ≡ go<B Equal a b
|
|
||||||
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
|
||||||
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) | inl (inr b<a) = exFalso (noIntegersBetweenXAndSuccX {2 *N binNatToN a} (2 *N binNatToN b) (lessRespectsMultiplicationLeft (binNatToN b) (binNatToN a) 2 b<a (le 1 refl)) 2a<2b+1)
|
|
||||||
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) | inr a=b rewrite a=b | canonicalFirst a b FirstLess | canonicalSecond (canonical a) b FirstLess | transitivity (equalityCommutative (binToBin a)) (transitivity (applyEquality NToBinNat a=b) (binToBin b)) = equalityCommutative (lemma1 (canonical b))
|
|
||||||
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
|
||||||
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) | inl (inl a<b) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) 2b+1<2a (TotalOrder.<Transitive (ℕTotalOrder) (lessRespectsMultiplicationLeft (binNatToN a) (binNatToN b) 2 a<b (le 1 refl)) (le zero refl))))
|
|
||||||
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) | inl (inr b<a) = equalityCommutative (equalToFirstGreater FirstLess a b (equalityCommutative indHyp))
|
|
||||||
where
|
|
||||||
t : a <BInherited b ≡ FirstGreater
|
|
||||||
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
|
||||||
t | inl (inl x) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) x b<a))
|
|
||||||
t | inl (inr x) = refl
|
|
||||||
t | inr x rewrite x = exFalso (TotalOrder.irreflexive (ℕTotalOrder) b<a)
|
|
||||||
indHyp : FirstGreater ≡ go<B Equal a b
|
|
||||||
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
|
||||||
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) | inr a=b rewrite a=b = exFalso (succNotLess 2b+1<2a)
|
|
||||||
<BIsInherited (zero :: a) (one :: b) | inr 2a=2b+1 = exFalso (parity (binNatToN b) (binNatToN a) (equalityCommutative 2a=2b+1))
|
|
||||||
<BIsInherited (one :: a) (zero :: b) with TotalOrder.totality ℕTotalOrder (binNatToN (one :: a)) (binNatToN (zero :: b))
|
|
||||||
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
|
||||||
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) | inl (inl a<b) = equalityCommutative (equalToFirstLess FirstGreater a b (equalityCommutative indHyp))
|
|
||||||
where
|
where
|
||||||
t : a <BInherited b ≡ FirstLess
|
t : a <BInherited b ≡ FirstLess
|
||||||
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||||
t | inl (inr x) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) x a<b))
|
|
||||||
t | inl (inl x) = refl
|
t | inl (inl x) = refl
|
||||||
|
t | inl (inr x) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) x a<b))
|
||||||
t | inr x rewrite x = exFalso (TotalOrder.irreflexive (ℕTotalOrder) a<b)
|
t | inr x rewrite x = exFalso (TotalOrder.irreflexive (ℕTotalOrder) a<b)
|
||||||
indHyp : FirstLess ≡ go<B Equal a b
|
indHyp : FirstLess ≡ go<B Equal a b
|
||||||
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
||||||
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) | inl (inr b<a) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) 2a+1<2b (TotalOrder.<Transitive (ℕTotalOrder) (lessRespectsMultiplicationLeft (binNatToN b) (binNatToN a) 2 b<a (le 1 refl)) (le zero refl))))
|
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) | inl (inr b<a) = exFalso (noIntegersBetweenXAndSuccX {2 *N binNatToN a} (2 *N binNatToN b) (lessRespectsMultiplicationLeft (binNatToN b) (binNatToN a) 2 b<a (le 1 refl)) 2a<2b+1)
|
||||||
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) | inr a=b rewrite a=b = exFalso (succNotLess 2a+1<2b)
|
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) | inr a=b rewrite a=b | canonicalFirst a b FirstLess | canonicalSecond (canonical a) b FirstLess | transitivity (equalityCommutative (binToBin a)) (transitivity (applyEquality NToBinNat a=b) (binToBin b)) = equalityCommutative (lemma1 (canonical b))
|
||||||
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||||
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) | inl (inl a<b) = exFalso (noIntegersBetweenXAndSuccX {2 *N binNatToN b} (2 *N binNatToN a) (lessRespectsMultiplicationLeft (binNatToN a) (binNatToN b) 2 a<b (le 1 refl)) 2b<2a+1)
|
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) | inl (inl a<b) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) 2b+1<2a (TotalOrder.<Transitive (ℕTotalOrder) (lessRespectsMultiplicationLeft (binNatToN a) (binNatToN b) 2 a<b (le 1 refl)) (le zero refl))))
|
||||||
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) | inl (inr b<a) = equalityCommutative (equalToFirstGreater FirstGreater a b (equalityCommutative indHyp))
|
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) | inl (inr b<a) = equalityCommutative (equalToFirstGreater FirstLess a b (equalityCommutative indHyp))
|
||||||
where
|
where
|
||||||
t : a <BInherited b ≡ FirstGreater
|
t : a <BInherited b ≡ FirstGreater
|
||||||
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||||
t | inl (inl x) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) x b<a))
|
t | inl (inl x) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) x b<a))
|
||||||
t | inl (inr x) = refl
|
t | inl (inr x) = refl
|
||||||
t | inr x rewrite x = exFalso (TotalOrder.irreflexive (ℕTotalOrder) b<a)
|
t | inr x rewrite x = exFalso (TotalOrder.irreflexive (ℕTotalOrder) b<a)
|
||||||
indHyp : FirstGreater ≡ go<B Equal a b
|
indHyp : FirstGreater ≡ go<B Equal a b
|
||||||
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
||||||
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) | inr a=b rewrite a=b | canonicalFirst a b FirstGreater | canonicalSecond (canonical a) b FirstGreater | transitivity (equalityCommutative (binToBin a)) (transitivity (applyEquality NToBinNat a=b) (binToBin b)) = equalityCommutative (lemma1 (canonical b))
|
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) | inr a=b rewrite a=b = exFalso (succNotLess 2b+1<2a)
|
||||||
<BIsInherited (one :: a) (zero :: b) | inr x = exFalso (parity (binNatToN a) (binNatToN b) x)
|
<BIsInherited (zero :: a) (one :: b) | inr 2a=2b+1 = exFalso (parity (binNatToN b) (binNatToN a) (equalityCommutative 2a=2b+1))
|
||||||
<BIsInherited (one :: a) (one :: b) = transitivity (chopDouble a b one) (<BIsInherited a b)
|
<BIsInherited (one :: a) (zero :: b) with TotalOrder.totality ℕTotalOrder (binNatToN (one :: a)) (binNatToN (zero :: b))
|
||||||
|
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||||
|
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) | inl (inl a<b) = equalityCommutative (equalToFirstLess FirstGreater a b (equalityCommutative indHyp))
|
||||||
|
where
|
||||||
|
t : a <BInherited b ≡ FirstLess
|
||||||
|
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||||
|
t | inl (inr x) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) x a<b))
|
||||||
|
t | inl (inl x) = refl
|
||||||
|
t | inr x rewrite x = exFalso (TotalOrder.irreflexive (ℕTotalOrder) a<b)
|
||||||
|
indHyp : FirstLess ≡ go<B Equal a b
|
||||||
|
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
||||||
|
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) | inl (inr b<a) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) 2a+1<2b (TotalOrder.<Transitive (ℕTotalOrder) (lessRespectsMultiplicationLeft (binNatToN b) (binNatToN a) 2 b<a (le 1 refl)) (le zero refl))))
|
||||||
|
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) | inr a=b rewrite a=b = exFalso (succNotLess 2a+1<2b)
|
||||||
|
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||||
|
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) | inl (inl a<b) = exFalso (noIntegersBetweenXAndSuccX {2 *N binNatToN b} (2 *N binNatToN a) (lessRespectsMultiplicationLeft (binNatToN a) (binNatToN b) 2 a<b (le 1 refl)) 2b<2a+1)
|
||||||
|
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) | inl (inr b<a) = equalityCommutative (equalToFirstGreater FirstGreater a b (equalityCommutative indHyp))
|
||||||
|
where
|
||||||
|
t : a <BInherited b ≡ FirstGreater
|
||||||
|
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||||
|
t | inl (inl x) = exFalso (TotalOrder.irreflexive (ℕTotalOrder) (TotalOrder.<Transitive (ℕTotalOrder) x b<a))
|
||||||
|
t | inl (inr x) = refl
|
||||||
|
t | inr x rewrite x = exFalso (TotalOrder.irreflexive (ℕTotalOrder) b<a)
|
||||||
|
indHyp : FirstGreater ≡ go<B Equal a b
|
||||||
|
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
||||||
|
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) | inr a=b rewrite a=b | canonicalFirst a b FirstGreater | canonicalSecond (canonical a) b FirstGreater | transitivity (equalityCommutative (binToBin a)) (transitivity (applyEquality NToBinNat a=b) (binToBin b)) = equalityCommutative (lemma1 (canonical b))
|
||||||
|
<BIsInherited (one :: a) (zero :: b) | inr x = exFalso (parity (binNatToN a) (binNatToN b) x)
|
||||||
|
<BIsInherited (one :: a) (one :: b) = transitivity (chopDouble a b one) (<BIsInherited a b)
|
||||||
|
@@ -15,6 +15,7 @@ open import Maybe
|
|||||||
|
|
||||||
module Numbers.BinaryNaturals.Subtraction where
|
module Numbers.BinaryNaturals.Subtraction where
|
||||||
|
|
||||||
|
private
|
||||||
aMinusAGo : (a : BinNat) → mapMaybe canonical (go zero a a) ≡ yes []
|
aMinusAGo : (a : BinNat) → mapMaybe canonical (go zero a a) ≡ yes []
|
||||||
aMinusAGo [] = refl
|
aMinusAGo [] = refl
|
||||||
aMinusAGo (zero :: a) with aMinusAGo a
|
aMinusAGo (zero :: a) with aMinusAGo a
|
||||||
@@ -397,13 +398,14 @@ module Numbers.BinaryNaturals.Subtraction where
|
|||||||
bad' : one :: a :: as ≡ one :: [] → False
|
bad' : one :: a :: as ≡ one :: [] → False
|
||||||
bad' ()
|
bad' ()
|
||||||
|
|
||||||
doublingLemma : (y : BinNat) → NToBinNat (2 *N binNatToN y) ≡ canonical (zero :: y)
|
doublingLemma : (y : BinNat) → NToBinNat (2 *N binNatToN y) ≡ canonical (zero :: y)
|
||||||
doublingLemma y with inspect (canonical y)
|
doublingLemma y with inspect (canonical y)
|
||||||
doublingLemma y | [] with≡ pr rewrite binNatToNZero' y pr | pr = refl
|
doublingLemma y | [] with≡ pr rewrite binNatToNZero' y pr | pr = refl
|
||||||
doublingLemma y | (a :: as) with≡ pr with inspect (binNatToN y)
|
doublingLemma y | (a :: as) with≡ pr with inspect (binNatToN y)
|
||||||
doublingLemma y | (a :: as) with≡ pr | zero with≡ pr2 rewrite binNatToNZero y pr2 = exFalso (nonEmptyNotEmpty (equalityCommutative pr))
|
doublingLemma y | (a :: as) with≡ pr | zero with≡ pr2 rewrite binNatToNZero y pr2 = exFalso (nonEmptyNotEmpty (equalityCommutative pr))
|
||||||
doublingLemma y | (a :: as) with≡ pr | succ bl with≡ pr2 rewrite pr | pr2 | doubleIsBitShift' bl | equalityCommutative pr = applyEquality (zero ::_) (equalityCommutative (transitivity (equalityCommutative (binToBin y)) (applyEquality NToBinNat pr2)))
|
doublingLemma y | (a :: as) with≡ pr | succ bl with≡ pr2 rewrite pr | pr2 | doubleIsBitShift' bl | equalityCommutative pr = applyEquality (zero ::_) (equalityCommutative (transitivity (equalityCommutative (binToBin y)) (applyEquality NToBinNat pr2)))
|
||||||
|
|
||||||
|
private
|
||||||
doubling : (a : ℕ) {y : BinNat} → (NToBinNat a ≡ zero :: y) → binNatToN y +N (binNatToN y +N 0) ≡ a
|
doubling : (a : ℕ) {y : BinNat} → (NToBinNat a ≡ zero :: y) → binNatToN y +N (binNatToN y +N 0) ≡ a
|
||||||
doubling a {y} pr = NToBinNatInj (binNatToN y +N (binNatToN y +N zero)) a (transitivity (transitivity (equalityCommutative (NToBinNatIsCanonical (binNatToN y +N (binNatToN y +N zero)))) (doublingLemma y)) (applyEquality canonical (equalityCommutative pr)))
|
doubling a {y} pr = NToBinNatInj (binNatToN y +N (binNatToN y +N zero)) a (transitivity (transitivity (equalityCommutative (NToBinNatIsCanonical (binNatToN y +N (binNatToN y +N zero)))) (doublingLemma y)) (applyEquality canonical (equalityCommutative pr)))
|
||||||
|
|
||||||
|
@@ -7,77 +7,77 @@ open import Maybe
|
|||||||
|
|
||||||
module Numbers.BinaryNaturals.SubtractionGo where
|
module Numbers.BinaryNaturals.SubtractionGo where
|
||||||
|
|
||||||
go : Bit → BinNat → BinNat → Maybe BinNat
|
go : Bit → BinNat → BinNat → Maybe BinNat
|
||||||
go zero [] [] = yes []
|
go zero [] [] = yes []
|
||||||
go one [] [] = no
|
go one [] [] = no
|
||||||
go zero [] (zero :: b) = go zero [] b
|
go zero [] (zero :: b) = go zero [] b
|
||||||
go zero [] (one :: b) = no
|
go zero [] (one :: b) = no
|
||||||
go one [] (x :: b) = no
|
go one [] (x :: b) = no
|
||||||
go zero (zero :: a) [] = yes (zero :: a)
|
go zero (zero :: a) [] = yes (zero :: a)
|
||||||
go one (zero :: a) [] = mapMaybe (one ::_) (go one a [])
|
go one (zero :: a) [] = mapMaybe (one ::_) (go one a [])
|
||||||
go zero (zero :: a) (zero :: b) = mapMaybe (zero ::_) (go zero a b)
|
go zero (zero :: a) (zero :: b) = mapMaybe (zero ::_) (go zero a b)
|
||||||
go one (zero :: a) (zero :: b) = mapMaybe (one ::_) (go one a b)
|
go one (zero :: a) (zero :: b) = mapMaybe (one ::_) (go one a b)
|
||||||
go zero (zero :: a) (one :: b) = mapMaybe (one ::_) (go one a b)
|
go zero (zero :: a) (one :: b) = mapMaybe (one ::_) (go one a b)
|
||||||
go one (zero :: a) (one :: b) = mapMaybe (zero ::_) (go one a b)
|
go one (zero :: a) (one :: b) = mapMaybe (zero ::_) (go one a b)
|
||||||
go zero (one :: a) [] = yes (one :: a)
|
go zero (one :: a) [] = yes (one :: a)
|
||||||
go zero (one :: a) (zero :: b) = mapMaybe (one ::_) (go zero a b)
|
go zero (one :: a) (zero :: b) = mapMaybe (one ::_) (go zero a b)
|
||||||
go zero (one :: a) (one :: b) = mapMaybe (zero ::_) (go zero a b)
|
go zero (one :: a) (one :: b) = mapMaybe (zero ::_) (go zero a b)
|
||||||
go one (one :: a) [] = yes (zero :: a)
|
go one (one :: a) [] = yes (zero :: a)
|
||||||
go one (one :: a) (zero :: b) = mapMaybe (zero ::_) (go zero a b)
|
go one (one :: a) (zero :: b) = mapMaybe (zero ::_) (go zero a b)
|
||||||
go one (one :: a) (one :: b) = mapMaybe (one ::_) (go one a b)
|
go one (one :: a) (one :: b) = mapMaybe (one ::_) (go one a b)
|
||||||
|
|
||||||
_-B_ : BinNat → BinNat → Maybe BinNat
|
_-B_ : BinNat → BinNat → Maybe BinNat
|
||||||
a -B b = go zero a b
|
a -B b = go zero a b
|
||||||
|
|
||||||
goEmpty : (a : BinNat) → go zero a [] ≡ yes a
|
goEmpty : (a : BinNat) → go zero a [] ≡ yes a
|
||||||
goEmpty [] = refl
|
goEmpty [] = refl
|
||||||
goEmpty (zero :: a) = refl
|
goEmpty (zero :: a) = refl
|
||||||
goEmpty (one :: a) = refl
|
goEmpty (one :: a) = refl
|
||||||
|
|
||||||
goOneSelf : (a : BinNat) → go one a a ≡ no
|
goOneSelf : (a : BinNat) → go one a a ≡ no
|
||||||
goOneSelf [] = refl
|
goOneSelf [] = refl
|
||||||
goOneSelf (zero :: a) rewrite goOneSelf a = refl
|
goOneSelf (zero :: a) rewrite goOneSelf a = refl
|
||||||
goOneSelf (one :: a) rewrite goOneSelf a = refl
|
goOneSelf (one :: a) rewrite goOneSelf a = refl
|
||||||
|
|
||||||
goOneEmpty : (b : BinNat) {t : BinNat} → go one [] b ≡ yes t → False
|
goOneEmpty : (b : BinNat) {t : BinNat} → go one [] b ≡ yes t → False
|
||||||
goOneEmpty [] {t} ()
|
goOneEmpty [] {t} ()
|
||||||
goOneEmpty (x :: b) {t} ()
|
goOneEmpty (x :: b) {t} ()
|
||||||
|
|
||||||
goOneEmpty' : (b : BinNat) → go one [] b ≡ no
|
goOneEmpty' : (b : BinNat) → go one [] b ≡ no
|
||||||
goOneEmpty' b with inspect (go one [] b)
|
goOneEmpty' b with inspect (go one [] b)
|
||||||
goOneEmpty' b | no with≡ x = x
|
goOneEmpty' b | no with≡ x = x
|
||||||
goOneEmpty' b | yes x₁ with≡ x = exFalso (goOneEmpty b x)
|
goOneEmpty' b | yes x₁ with≡ x = exFalso (goOneEmpty b x)
|
||||||
|
|
||||||
goZeroEmpty : (b : BinNat) {t : BinNat} → go zero [] b ≡ yes t → canonical b ≡ []
|
goZeroEmpty : (b : BinNat) {t : BinNat} → go zero [] b ≡ yes t → canonical b ≡ []
|
||||||
goZeroEmpty [] {t} = λ _ → refl
|
goZeroEmpty [] {t} = λ _ → refl
|
||||||
goZeroEmpty (zero :: b) {t} pr with inspect (canonical b)
|
goZeroEmpty (zero :: b) {t} pr with inspect (canonical b)
|
||||||
goZeroEmpty (zero :: b) {t} pr | [] with≡ pr2 rewrite pr2 = refl
|
goZeroEmpty (zero :: b) {t} pr | [] with≡ pr2 rewrite pr2 = refl
|
||||||
goZeroEmpty (zero :: b) {t} pr | (x :: r) with≡ pr2 with goZeroEmpty b pr
|
goZeroEmpty (zero :: b) {t} pr | (x :: r) with≡ pr2 with goZeroEmpty b pr
|
||||||
... | u = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative pr2) u))
|
... | u = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative pr2) u))
|
||||||
|
|
||||||
goZeroEmpty' : (b : BinNat) {t : BinNat} → go zero [] b ≡ yes t → canonical t ≡ []
|
goZeroEmpty' : (b : BinNat) {t : BinNat} → go zero [] b ≡ yes t → canonical t ≡ []
|
||||||
goZeroEmpty' [] {[]} pr = refl
|
goZeroEmpty' [] {[]} pr = refl
|
||||||
goZeroEmpty' (x :: b) {[]} pr = refl
|
goZeroEmpty' (x :: b) {[]} pr = refl
|
||||||
goZeroEmpty' (zero :: b) {x₁ :: t} pr = goZeroEmpty' b pr
|
goZeroEmpty' (zero :: b) {x₁ :: t} pr = goZeroEmpty' b pr
|
||||||
|
|
||||||
goZeroIncr : (b : BinNat) → go zero [] (incr b) ≡ no
|
goZeroIncr : (b : BinNat) → go zero [] (incr b) ≡ no
|
||||||
goZeroIncr [] = refl
|
goZeroIncr [] = refl
|
||||||
goZeroIncr (zero :: b) = refl
|
goZeroIncr (zero :: b) = refl
|
||||||
goZeroIncr (one :: b) = goZeroIncr b
|
goZeroIncr (one :: b) = goZeroIncr b
|
||||||
|
|
||||||
goPreservesCanonicalRightEmpty : (b : BinNat) → go zero [] (canonical b) ≡ go zero [] b
|
goPreservesCanonicalRightEmpty : (b : BinNat) → go zero [] (canonical b) ≡ go zero [] b
|
||||||
goPreservesCanonicalRightEmpty [] = refl
|
goPreservesCanonicalRightEmpty [] = refl
|
||||||
goPreservesCanonicalRightEmpty (zero :: b) with inspect (canonical b)
|
goPreservesCanonicalRightEmpty (zero :: b) with inspect (canonical b)
|
||||||
goPreservesCanonicalRightEmpty (zero :: b) | [] with≡ x with goPreservesCanonicalRightEmpty b
|
goPreservesCanonicalRightEmpty (zero :: b) | [] with≡ x with goPreservesCanonicalRightEmpty b
|
||||||
... | pr2 rewrite x = pr2
|
... | pr2 rewrite x = pr2
|
||||||
goPreservesCanonicalRightEmpty (zero :: b) | (x₁ :: y) with≡ x with goPreservesCanonicalRightEmpty b
|
goPreservesCanonicalRightEmpty (zero :: b) | (x₁ :: y) with≡ x with goPreservesCanonicalRightEmpty b
|
||||||
... | pr2 rewrite x = pr2
|
... | pr2 rewrite x = pr2
|
||||||
goPreservesCanonicalRightEmpty (one :: b) = refl
|
goPreservesCanonicalRightEmpty (one :: b) = refl
|
||||||
|
|
||||||
goZero : (b : BinNat) {t : BinNat} → mapMaybe canonical (go zero [] b) ≡ yes t → t ≡ []
|
goZero : (b : BinNat) {t : BinNat} → mapMaybe canonical (go zero [] b) ≡ yes t → t ≡ []
|
||||||
goZero b {[]} pr = refl
|
goZero b {[]} pr = refl
|
||||||
goZero b {x :: t} pr with inspect (go zero [] b)
|
goZero b {x :: t} pr with inspect (go zero [] b)
|
||||||
goZero b {x :: t} pr | no with≡ pr2 rewrite pr2 = exFalso (noNotYes pr)
|
goZero b {x :: t} pr | no with≡ pr2 rewrite pr2 = exFalso (noNotYes pr)
|
||||||
goZero b {x :: t} pr | yes x₁ with≡ pr2 with goZeroEmpty b pr2
|
goZero b {x :: t} pr | yes x₁ with≡ pr2 with goZeroEmpty b pr2
|
||||||
... | u with applyEquality (mapMaybe canonical) (goPreservesCanonicalRightEmpty b)
|
... | u with applyEquality (mapMaybe canonical) (goPreservesCanonicalRightEmpty b)
|
||||||
... | bl rewrite u | pr = exFalso (nonEmptyNotEmpty (equalityCommutative (yesInjective bl)))
|
... | bl rewrite u | pr = exFalso (nonEmptyNotEmpty (equalityCommutative (yesInjective bl)))
|
||||||
|
@@ -33,3 +33,5 @@ record Ring {n m} {A : Set n} (S : Setoid {n} {m} A) (_+_ : A → A → A) (_*_
|
|||||||
timesZero' {a} = transitive *Commutative timesZero
|
timesZero' {a} = transitive *Commutative timesZero
|
||||||
*DistributesOver+' : {a b c : A} → (a + b) * c ∼ (a * c) + (b * c)
|
*DistributesOver+' : {a b c : A} → (a + b) * c ∼ (a * c) + (b * c)
|
||||||
*DistributesOver+' = transitive *Commutative (transitive *DistributesOver+ (+WellDefined *Commutative *Commutative))
|
*DistributesOver+' = transitive *Commutative (transitive *DistributesOver+ (+WellDefined *Commutative *Commutative))
|
||||||
|
*Associative' : {a b c : A} → ((a * b) * c) ∼ (a * (b * c))
|
||||||
|
*Associative' = symmetric *Associative
|
||||||
|
Reference in New Issue
Block a user