mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-06 12:28:39 +00:00
Hide more stuff (#124)
This commit is contained in:
@@ -26,3 +26,8 @@ boolAnd BoolFalse y = BoolFalse
|
||||
boolOr : Bool → Bool → Bool
|
||||
boolOr BoolTrue y = BoolTrue
|
||||
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.Lemmas
|
||||
open import Fields.FieldOfFractions.Order
|
||||
open import Fields.FieldOfFractions.Order
|
||||
--open import Fields.FieldOfFractions.Archimedean
|
||||
|
||||
open import Rings.Definition
|
||||
|
@@ -18,24 +18,9 @@ fieldOfFractionsRing : Ring fieldOfFractionsSetoid fieldOfFractionsPlus fieldOfF
|
||||
Ring.additiveGroup fieldOfFractionsRing = fieldOfFractionsGroup
|
||||
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.groupIsAbelian fieldOfFractionsRing {record { num = a ; denom = b }} {record { num = c ; denom = d }} = need
|
||||
where
|
||||
open Setoid S
|
||||
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.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)))
|
||||
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)
|
||||
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))
|
||||
Ring.*DistributesOver+ fieldOfFractionsRing {record { num = a ; denom = b }} {record { num = c ; denom = d }} {record { num = e ; denom = f }} = need
|
||||
where
|
||||
open Setoid S
|
||||
|
@@ -12,17 +12,13 @@ open import Semirings.Definition
|
||||
open import Functions.Definition
|
||||
open import Groups.Isomorphisms.Definition
|
||||
open import Boolean.Definition
|
||||
open import Boolean.Lemmas
|
||||
|
||||
module Groups.FreeGroup.Parity {a : _} {A : Set a} (decA : DecidableSet A) where
|
||||
|
||||
open import Groups.FreeGroup.Word 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
|
||||
Group.+WellDefined C2 refl refl = refl
|
||||
Group.0G C2 = BoolFalse
|
||||
@@ -39,14 +35,6 @@ Group.invLeft C2 {BoolFalse} = refl
|
||||
Group.invRight C2 {BoolTrue} = 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 empty = BoolFalse
|
||||
parity x (prependLetter (ofLetter y) w _) with decA x y
|
||||
@@ -56,6 +44,7 @@ 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 _) | inr _ = parity x w
|
||||
|
||||
private
|
||||
parityPrepend : (a : A) (w : ReducedWord) (l : A) → ((a ≡ l) → False) → parity a (prepend w (ofLetter l)) ≡ parity a w
|
||||
parityPrepend a empty l notEq with decA a l
|
||||
parityPrepend a empty l notEq | inl x = exFalso (notEq x)
|
||||
@@ -144,19 +133,13 @@ parityPrepend''' a (prependLetter (ofInv l) w x) | inl _ with decA a l
|
||||
... | inr a!=l = 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
|
||||
notXor BoolTrue BoolTrue = refl
|
||||
notXor BoolTrue BoolFalse = refl
|
||||
notXor BoolFalse BoolTrue = refl
|
||||
notXor BoolFalse BoolFalse = refl
|
||||
|
||||
parityHomIsHom : (a : A) → (x y : ReducedWord) → parity a (_+W_ x y) ≡ xor (parity a x) (parity a y)
|
||||
parityHomIsHom a empty y = refl
|
||||
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 .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 (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 .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 (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)
|
||||
|
@@ -19,6 +19,7 @@ open import Groups.FreeProduct.Setoid decidableIndex decidableGroups G
|
||||
open import Groups.FreeProduct.Addition decidableIndex decidableGroups G
|
||||
open import Groups.FreeProduct.Group decidableIndex decidableGroups G
|
||||
|
||||
private
|
||||
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} (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
|
||||
@@ -27,6 +28,7 @@ universalPropertyFunction : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ :
|
||||
universalPropertyFunction H fs homs empty = Group.0G H
|
||||
universalPropertyFunction H fs homs (nonempty i x) = universalPropertyFunction' H fs homs x
|
||||
|
||||
private
|
||||
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))
|
||||
upWellDefined' H fs homs (ofEmpty m g nonZero) (ofEmpty n g₁ nonZero₁) eq with decidableIndex m n
|
||||
... | inl refl = GroupHom.wellDefined (homs m) eq
|
||||
@@ -37,6 +39,7 @@ upWellDefined : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+_ : C → C →
|
||||
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
|
||||
|
||||
private
|
||||
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 empty g pr = Equivalence.symmetric (Setoid.eq T) (Group.identRight H)
|
||||
upPrepend {T = T} H fs homs {j} (nonempty i (ofEmpty .i h nonZero)) g pr with decidableIndex j i
|
||||
@@ -56,6 +59,7 @@ upPrepend {T = T} H fs homs {j} (nonempty k (prependLetter .k h nonZero y _)) g
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
private
|
||||
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 _ g nonZero) empty = Equivalence.symmetric (Setoid.eq T) (Group.identRight H)
|
||||
upHom {T = T} H fs homs (ofEmpty j g nonZero) (nonempty i (ofEmpty .i h nonZero1)) with decidableIndex j i
|
||||
@@ -86,6 +90,7 @@ 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 {T = T} H fs homs g nz = Equivalence.reflexive (Setoid.eq T)
|
||||
|
||||
private
|
||||
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)
|
||||
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))))
|
||||
where
|
||||
|
@@ -18,6 +18,7 @@ _*B_ : BinNat → BinNat → BinNat
|
||||
(zero :: a) *B b = zero :: (a *B b)
|
||||
(one :: a) *B b = (zero :: (a *B b)) +B b
|
||||
|
||||
private
|
||||
contr : {a : _} {A : Set a} {l1 l2 : List A} → {x : A} → l1 ≡ [] → l1 ≡ x :: l2 → False
|
||||
contr {l1 = []} p1 ()
|
||||
contr {l1 = x :: l1} () p2
|
||||
|
@@ -16,6 +16,7 @@ module Numbers.BinaryNaturals.Order where
|
||||
FirstLess : Compare
|
||||
FirstGreater : Compare
|
||||
|
||||
private
|
||||
badCompare : Equal ≡ FirstLess → False
|
||||
badCompare ()
|
||||
|
||||
@@ -31,6 +32,7 @@ module Numbers.BinaryNaturals.Order where
|
||||
(a <BInherited b) | inl (inr x) = FirstGreater
|
||||
(a <BInherited b) | inr x = Equal
|
||||
|
||||
private
|
||||
go<B : Compare → BinNat → BinNat → Compare
|
||||
go<B Equal [] [] = Equal
|
||||
go<B Equal [] (zero :: b) = go<B Equal [] b
|
||||
@@ -61,6 +63,7 @@ module Numbers.BinaryNaturals.Order where
|
||||
_<B_ : BinNat → BinNat → Compare
|
||||
a <B b = go<B Equal a b
|
||||
|
||||
private
|
||||
lemma1 : {s : Compare} → (n : BinNat) → go<B s n n ≡ s
|
||||
lemma1 {Equal} [] = refl
|
||||
lemma1 {Equal} (zero :: n) = lemma1 n
|
||||
@@ -304,7 +307,7 @@ 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 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 i | inr a=b | inr x = refl
|
||||
chopDouble a b _ | inr a=b | inr x = refl
|
||||
|
||||
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)))
|
||||
|
@@ -15,6 +15,7 @@ open import Maybe
|
||||
|
||||
module Numbers.BinaryNaturals.Subtraction where
|
||||
|
||||
private
|
||||
aMinusAGo : (a : BinNat) → mapMaybe canonical (go zero a a) ≡ yes []
|
||||
aMinusAGo [] = refl
|
||||
aMinusAGo (zero :: a) with aMinusAGo a
|
||||
@@ -404,6 +405,7 @@ module Numbers.BinaryNaturals.Subtraction where
|
||||
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)))
|
||||
|
||||
private
|
||||
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)))
|
||||
|
||||
|
@@ -33,3 +33,5 @@ record Ring {n m} {A : Set n} (S : Setoid {n} {m} A) (_+_ : A → A → A) (_*_
|
||||
timesZero' {a} = transitive *Commutative timesZero
|
||||
*DistributesOver+' : {a b c : A} → (a + b) * c ∼ (a * c) + (b * c)
|
||||
*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