mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-05 11:58:41 +00:00
Rem unused opens in Safe (#93)
This commit is contained in:
@@ -2,10 +2,6 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Vectors
|
||||
open import Semirings.Definition
|
||||
|
||||
module Categories.Definition where
|
||||
|
||||
|
@@ -1,11 +1,6 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Vectors
|
||||
open import Semirings.Definition
|
||||
open import Categories.Definition
|
||||
|
||||
module Categories.Dual.Definition where
|
||||
|
@@ -2,10 +2,6 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Vectors
|
||||
open import Semirings.Definition
|
||||
open import Categories.Definition
|
||||
|
||||
module Categories.Functor.Definition where
|
||||
|
@@ -1,11 +1,6 @@
|
||||
{-# OPTIONS --warning=error --without-K --safe #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Vectors
|
||||
open import Semirings.Definition
|
||||
open import Categories.Definition
|
||||
open import Categories.Functor.Definition
|
||||
|
||||
|
@@ -2,7 +2,6 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module ClassicalLogic.ClassicalFive where
|
||||
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Decidable.Sets where
|
||||
|
||||
|
@@ -1,14 +1,11 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
|
@@ -1,13 +1,11 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
@@ -37,7 +35,6 @@ open import Rings.Orders.Total.Lemmas order
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Addition order F charNot2
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
open import Fields.CauchyCompletion.Comparison order F charNot2
|
||||
|
||||
abstract
|
||||
|
@@ -3,14 +3,11 @@
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Fields.Orders.Total.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
|
@@ -3,11 +3,9 @@
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
|
@@ -1,15 +1,12 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Groups
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
|
@@ -1,13 +1,11 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
@@ -18,7 +16,6 @@ open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Semirings.Definition
|
||||
|
||||
module Fields.CauchyCompletion.Multiplication {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
|
||||
@@ -37,8 +34,6 @@ open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
open import Fields.CauchyCompletion.Comparison order F charNot2
|
||||
open import Fields.CauchyCompletion.Addition order F charNot2
|
||||
open import Fields.CauchyCompletion.Approximation order F charNot2
|
||||
|
||||
0!=1 : {e : A} → (0G < e) → 0R ∼ 1R → False
|
||||
|
@@ -1,13 +1,10 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
@@ -17,7 +14,6 @@ open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
|
||||
module Fields.CauchyCompletion.Ring {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
|
||||
@@ -29,7 +25,6 @@ open PartiallyOrderedRing pRing
|
||||
open Field F
|
||||
open Group (Ring.additiveGroup R)
|
||||
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Multiplication order F charNot2
|
||||
|
@@ -7,7 +7,6 @@ open import Rings.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
|
@@ -1,18 +1,12 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.FieldOfFractions.Addition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) where
|
||||
|
||||
|
@@ -1,24 +1,15 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.FieldOfFractions.Field {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) where
|
||||
|
||||
open import Fields.FieldOfFractions.Setoid I
|
||||
open import Fields.FieldOfFractions.Addition I
|
||||
open import Fields.FieldOfFractions.Multiplication I
|
||||
open import Fields.FieldOfFractions.Ring I
|
||||
|
||||
fieldOfFractions : Field fieldOfFractionsRing
|
||||
|
@@ -1,18 +1,12 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.FieldOfFractions.Group {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) where
|
||||
|
||||
|
@@ -1,28 +1,19 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.FieldOfFractions.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) where
|
||||
|
||||
open import Fields.FieldOfFractions.Setoid I
|
||||
open import Fields.FieldOfFractions.Addition I
|
||||
open import Fields.FieldOfFractions.Multiplication I
|
||||
open import Fields.FieldOfFractions.Ring I
|
||||
open import Fields.FieldOfFractions.Field I
|
||||
|
||||
embedIntoFieldOfFractions : A → fieldOfFractionsSet
|
||||
embedIntoFieldOfFractions a = a ,, (Ring.1R R , IntegralDomain.nontrivial I)
|
||||
|
@@ -1,18 +1,11 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.FieldOfFractions.Multiplication {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) where
|
||||
|
||||
|
@@ -1,22 +1,17 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Orders.Total.Lemmas
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.FieldOfFractions.Order {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} {_<_ : Rel {_} {c} A} {pOrder : SetoidPartialOrder S _<_} {pRing : PartiallyOrderedRing R pOrder} (I : IntegralDomain R) (order : TotallyOrderedRing pRing) where
|
||||
|
||||
|
@@ -1,18 +1,12 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.FieldOfFractions.Ring {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) where
|
||||
|
||||
|
@@ -1,13 +1,7 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.IntegralDomains.Definition
|
||||
|
@@ -1,14 +1,9 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
|
@@ -1,18 +1,11 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Rings.IntegralDomains.Definition
|
||||
|
||||
module Fields.Lemmas {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (F : Field R) where
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
@@ -16,7 +15,6 @@ open import Sets.EquivalenceRelations
|
||||
open import Fields.Fields
|
||||
open import Fields.Orders.Total.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.Orders.Lemmas {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {_} {o} A} {R : Ring S _+_ _*_} {pOrder : SetoidPartialOrder S _<_} {F : Field R} {pRing : PartiallyOrderedRing R pOrder} (oF : TotallyOrderedField F pRing) where
|
||||
|
||||
|
@@ -1,15 +1,10 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Fields.Fields
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
@@ -17,7 +12,6 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
module Fields.Orders.Partial.Definition {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (F : Field R) where
|
||||
|
||||
open Ring R
|
||||
open import Fields.Lemmas F
|
||||
|
||||
record PartiallyOrderedField {p} {_<_ : Rel {_} {p} A} (pOrder : SetoidPartialOrder S _<_) : Set (lsuc (m ⊔ n ⊔ p)) where
|
||||
field
|
||||
|
@@ -1,16 +1,11 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Fields.Fields
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
@@ -18,7 +13,6 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
module Fields.Orders.Total.Definition {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (F : Field R) where
|
||||
|
||||
open Ring R
|
||||
open import Fields.Lemmas F
|
||||
|
||||
record TotallyOrderedField {p} {_<_ : Rel {_} {p} A} {pOrder : SetoidPartialOrder S _<_} (pRing : PartiallyOrderedRing R pOrder) : Set (lsuc (m ⊔ n ⊔ p)) where
|
||||
field
|
||||
|
@@ -1,12 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.Abelian.Definition where
|
||||
|
||||
|
@@ -1,12 +1,7 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Abelian.Definition
|
||||
|
||||
module Groups.Abelian.DirectSum {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+1_ : A → A → A} {_+2_ : B → B → B} {G1' : Group S _+1_} {G2' : Group T _+2_} (G1 : AbelianGroup G1') (G2 : AbelianGroup G2') where
|
||||
|
@@ -2,16 +2,11 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.DirectSum
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.DirectSum.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Isomorphisms.Definition
|
||||
|
||||
module Groups.Abelian.Lemmas where
|
||||
|
@@ -2,13 +2,8 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.SymmetricGroups.Definition
|
||||
open import Groups.Actions.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
@@ -1,14 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.Actions.Definition where
|
||||
|
||||
|
@@ -1,13 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
open import Groups.Actions.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
|
@@ -1,15 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Actions.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Actions.Definition
|
||||
|
@@ -1,11 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Subset
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Groups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
|
@@ -3,15 +3,9 @@
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Integers.Addition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
|
||||
|
@@ -3,16 +3,10 @@
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Integers.Addition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Cyclic.Definition
|
||||
open import Semirings.Definition
|
||||
|
@@ -1,10 +1,7 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
|
||||
module Groups.Definition where
|
||||
|
||||
|
@@ -2,11 +2,7 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.DirectSum.Definition {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} (G : Group S _·A_) (H : Group T _·B_) where
|
||||
|
||||
|
@@ -1,13 +1,10 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Sets.Cardinality.Finite.Definition
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.FiniteGroups.Definition where
|
||||
|
||||
|
@@ -4,8 +4,6 @@ open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring -- for length
|
||||
open import Lists.Lists
|
||||
open import Functions
|
||||
open import Groups.SymmetricGroups.Definition
|
||||
open import Setoids.Setoids
|
||||
--open import Groups.Actions
|
||||
|
||||
module Groups.FinitePermutations where
|
||||
|
@@ -1,23 +1,15 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Isomorphisms.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Subgroups.Normal.Definition
|
||||
open import Groups.Subgroups.Normal.Lemmas
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Abelian.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.FirstIsomorphismTheorem {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A → A → A} {_+H_ : B → B → B} {G : Group S _+G_} {H : Group T _+H_} {f : A → B} (fHom : GroupHom G H f) where
|
||||
|
||||
|
@@ -2,14 +2,8 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
|
||||
module Groups.Groups where
|
||||
|
||||
|
@@ -1,12 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.Homomorphisms.Definition where
|
||||
|
||||
|
@@ -1,14 +1,10 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Lemmas
|
||||
|
||||
module Groups.Homomorphisms.Examples where
|
||||
|
||||
|
@@ -1,21 +1,14 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Subset
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Isomorphisms.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.QuotientGroup.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
|
@@ -1,25 +1,14 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Isomorphisms.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Subgroups.Normal.Definition
|
||||
open import Groups.Subgroups.Normal.Lemmas
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.QuotientGroup.Definition
|
||||
open import Groups.Cosets
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Homomorphisms.Kernel {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A → A → A} {_+H_ : B → B → B} {G : Group S _+G_} {H : Group T _+H_} {f : A → B} (fHom : GroupHom G H f) where
|
||||
|
||||
|
@@ -1,10 +1,7 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Homomorphisms.Definition
|
||||
|
@@ -1,13 +1,9 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.Isomorphisms.Definition where
|
||||
|
||||
|
@@ -3,8 +3,6 @@
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Isomorphisms.Definition
|
||||
|
@@ -1,10 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
|
@@ -1,20 +1,13 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Vectors
|
||||
open import Lists.Lists
|
||||
open import Maybe
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Polynomials.Addition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} (G : Group S _+_) where
|
||||
|
||||
|
@@ -1,15 +1,10 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Vectors
|
||||
open import Lists.Lists
|
||||
open import Maybe
|
||||
|
||||
|
@@ -1,21 +1,11 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Integers.Definition
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Vectors
|
||||
open import Lists.Lists
|
||||
open import Maybe
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Polynomials.Examples where
|
||||
|
||||
|
@@ -1,20 +1,11 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Vectors
|
||||
open import Lists.Lists
|
||||
open import Maybe
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Polynomials.Group {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} (G : Group S _+_) where
|
||||
|
||||
|
@@ -1,17 +1,12 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Subgroups.Normal.Definition
|
||||
|
||||
module Groups.QuotientGroup.Definition {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_} {f : A → B} (fHom : GroupHom G H f) where
|
||||
|
||||
|
@@ -1,12 +1,8 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Lemmas
|
||||
|
@@ -2,11 +2,8 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
|
||||
module Groups.Subgroups.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} (G : Group S _+_) where
|
||||
|
||||
|
@@ -3,9 +3,6 @@
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
|
||||
module Groups.Subgroups.Examples {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} (G : Group S _+_) where
|
||||
|
@@ -1,19 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Isomorphisms.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Abelian.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
|
@@ -1,23 +1,10 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Isomorphisms.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.QuotientGroup.Definition
|
||||
open import Groups.Subgroups.Normal.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Subgroups.Normal.Examples {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} (G : Group S _+_) where
|
||||
|
||||
|
@@ -1,20 +1,13 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Isomorphisms.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.QuotientGroup.Definition
|
||||
open import Groups.Subgroups.Normal.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
@@ -2,10 +2,7 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Setoids.Functions.Extension
|
||||
|
@@ -1,14 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.QuotientGroup.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
|
@@ -3,10 +3,8 @@
|
||||
open import LogicalFormulae
|
||||
open import Maybe
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Vectors
|
||||
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
|
||||
module KeyValue.KeyValue {a b : _} (keys : Set a) (values : Set b) where
|
||||
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Definition
|
||||
open import Lists.Length
|
||||
open import Numbers.Naturals.Semiring
|
||||
|
@@ -1,7 +1,5 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
|
||||
module Lists.Definition {a : _} where
|
||||
|
||||
|
@@ -1,7 +1,5 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Definition
|
||||
|
||||
module Lists.Fold.Fold {a b : _} {A : Set a} {B : Set b} where
|
||||
|
@@ -1,10 +1,7 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Semiring -- for length
|
||||
open import Numbers.Naturals.Order
|
||||
open import Semirings.Definition
|
||||
open import Lists.Definition
|
||||
open import Lists.Fold.Fold
|
||||
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Semiring -- for length
|
||||
open import Numbers.Naturals.Order
|
||||
open import Semirings.Definition
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Definition
|
||||
open import Lists.Fold.Fold
|
||||
open import Lists.Concat
|
||||
|
@@ -1,11 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Semiring -- for length
|
||||
open import Numbers.Naturals.Order
|
||||
open import Semirings.Definition
|
||||
open import Maybe
|
||||
open import Lists.Definition
|
||||
open import Lists.Concat
|
||||
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Maybe where
|
||||
|
||||
|
@@ -1,15 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
@@ -1,19 +1,12 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Modules.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Modules.DirectSum {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+R_ : A → A → A} {_*R_ : A → A → A} (R : Ring S _+R_ _*R_) {m n o p : _} {M : Set m} {T : Setoid {m} {n} M} {_+_ : M → M → M} {G' : Group T _+_} {G : AbelianGroup G'} {_·1_ : A → M → M} {N : Set o} {U : Setoid {o} {p} N} {_+'_ : N → N → N} {H' : Group U _+'_} {H : AbelianGroup H'} {_·2_ : A → N → N} (M1 : Module R G _·1_) (M2 : Module R H _·2_) where
|
||||
|
||||
|
@@ -1,23 +1,15 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Modules.Definition
|
||||
open import Groups.Cyclic.Definition
|
||||
open import Groups.Cyclic.DefinitionLemmas
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Modules.Examples where
|
||||
|
||||
|
@@ -1,20 +1,13 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Abelian.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Modules.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Modules.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+R_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+R_ _*_} {m n : _} {M : Set m} {T : Setoid {m} {n} M} {_+_ : M → M → M} {G' : Group T _+_} {G : AbelianGroup G'} {_·_ : A → M → M} (mod : Module R G _·_) where
|
||||
|
||||
|
@@ -1,8 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Monoids.Definition where
|
||||
|
||||
|
@@ -1,12 +1,10 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Groups.Definition
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
open import Semirings.Definition
|
||||
open import Orders.Total.Definition
|
||||
|
@@ -1,12 +1,9 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Groups.Definition
|
||||
open import Semirings.Definition
|
||||
open import Orders.Total.Definition
|
||||
|
||||
|
@@ -1,11 +1,9 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
open import Numbers.BinaryNaturals.Addition
|
||||
open import Semirings.Definition
|
||||
|
@@ -1,13 +1,10 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import Orders.WellFounded.Induction
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Groups.Definition
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
open import Orders.Total.Definition
|
||||
open import Semirings.Definition
|
||||
|
@@ -9,7 +9,6 @@ open import Numbers.BinaryNaturals.Definition
|
||||
open import Numbers.BinaryNaturals.Addition
|
||||
open import Numbers.BinaryNaturals.SubtractionGo
|
||||
open import Numbers.BinaryNaturals.SubtractionGoPreservesCanonicalRight
|
||||
open import Numbers.BinaryNaturals.SubtractionGoPreservesCanonicalLeft
|
||||
open import Orders.Total.Definition
|
||||
open import Semirings.Definition
|
||||
open import Maybe
|
||||
|
@@ -2,9 +2,7 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
open import Semirings.Definition
|
||||
open import Maybe
|
||||
|
||||
module Numbers.BinaryNaturals.SubtractionGo where
|
||||
|
@@ -1,179 +0,0 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
open import Maybe
|
||||
open import Numbers.BinaryNaturals.SubtractionGo
|
||||
|
||||
module Numbers.BinaryNaturals.SubtractionGoPreservesCanonicalLeft where
|
||||
|
||||
goPreservesCanonicalLeftZero : (a b : BinNat) → mapMaybe canonical (go zero a b) ≡ mapMaybe canonical (go zero (canonical a) b)
|
||||
goPreservesCanonicalLeftOne : (a b : BinNat) → mapMaybe canonical (go one a b) ≡ mapMaybe canonical (go one (canonical a) b)
|
||||
|
||||
goPreservesCanonicalLeftZero [] b = refl
|
||||
goPreservesCanonicalLeftZero (zero :: a) [] with inspect (canonical a)
|
||||
goPreservesCanonicalLeftZero (zero :: a) [] | [] with≡ pr rewrite pr = refl
|
||||
goPreservesCanonicalLeftZero (zero :: a) [] | (x :: t) with≡ pr rewrite pr | transitivity (applyEquality canonical (equalityCommutative pr)) (transitivity (equalityCommutative (canonicalIdempotent a)) pr) = refl
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) with inspect (canonical a)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | [] with≡ pr with inspect (go zero a b)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | [] with≡ pr | no with≡ pr2 rewrite pr | pr2 = transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (transitivity (goPreservesCanonicalLeftZero a b) (applyEquality (λ i → mapMaybe canonical (go zero i b)) pr))
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | [] with≡ pr | yes x with≡ pr2 with inspect (canonical x)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | [] with≡ pr | yes x with≡ pr2 | [] with≡ pr3 rewrite pr | pr2 | pr3 = transitivity (equalityCommutative (transitivity (applyEquality (mapMaybe canonical) pr2) (applyEquality yes pr3))) (transitivity (goPreservesCanonicalLeftZero a b) (applyEquality (λ i → mapMaybe canonical (go zero i b)) pr))
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | [] with≡ pr | yes x with≡ pr2 | (y1 :: ys) with≡ pr3 = exFalso (nonEmptyNotEmpty (goZero b {y1 :: ys} t))
|
||||
where
|
||||
t : mapMaybe canonical (go zero [] b) ≡ yes (y1 :: ys)
|
||||
t = transitivity (transitivity (applyEquality (λ i → mapMaybe canonical (go zero i b)) (equalityCommutative pr)) (transitivity (equalityCommutative (goPreservesCanonicalLeftZero a b)) (applyEquality (mapMaybe canonical) pr2))) (applyEquality yes pr3)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | (x :: y) with≡ pr with inspect (go zero a b)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | (x :: y) with≡ pr | no with≡ pr2 with inspect (go zero (x :: y) b)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | (x :: y) with≡ pr | no with≡ pr2 | no with≡ pr3 rewrite pr | pr2 | pr3 = refl
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | (x :: y) with≡ pr | no with≡ pr2 | yes r with≡ pr3 = exFalso (noNotYes (transitivity (equalityCommutative (transitivity (applyEquality (λ i → mapMaybe canonical (go zero i b)) (equalityCommutative pr)) (transitivity (equalityCommutative (goPreservesCanonicalLeftZero a b)) (applyEquality (mapMaybe canonical) pr2)))) (applyEquality (mapMaybe canonical) pr3)))
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | (x :: y) with≡ pr | yes x1 with≡ pr2 with inspect (go zero (x :: y) b)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | (x :: y) with≡ pr | yes x1 with≡ pr2 | no with≡ pr3 rewrite equalityCommutative pr = exFalso (noNotYes {b = canonical x1} (transitivity (applyEquality (mapMaybe canonical) (equalityCommutative pr3)) (transitivity (equalityCommutative (goPreservesCanonicalLeftZero a b)) (applyEquality (mapMaybe canonical) pr2))))
|
||||
goPreservesCanonicalLeftZero (zero :: a) (zero :: b) | (x :: y) with≡ pr | yes x1 with≡ pr2 | yes x2 with≡ pr3 with yesInjective {x = canonical x1} {y = canonical x2} (transitivity (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (goPreservesCanonicalLeftZero a b)) (applyEquality (mapMaybe canonical) (transitivity (applyEquality (λ i → go zero i b) pr) pr3)))
|
||||
... | k rewrite pr | pr2 | pr3 | equalityCommutative pr = applyEquality yes t
|
||||
where
|
||||
t : canonical (zero :: x1) ≡ canonical (zero :: x2)
|
||||
t with inspect (canonical x1)
|
||||
t | [] with≡ pr rewrite pr | equalityCommutative k = refl
|
||||
t | (x :: bl) with≡ pr rewrite pr | equalityCommutative k = refl
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) with inspect (canonical a)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) | [] with≡ x with inspect (go one a b)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) | [] with≡ pr | no with≡ pr1 rewrite pr | pr1 = refl
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) | [] with≡ pr | yes x₂ with≡ pr1 with goPreservesCanonicalLeftOne a b
|
||||
... | bl rewrite pr1 | pr = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) (goOneEmpty' b))) (equalityCommutative bl)))
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) | (x₁ :: y) with≡ x with inspect (go one a b)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) | (r :: rs) with≡ pr1 | no with≡ pr2 with inspect (go one (r :: rs) b)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) | (r :: rs) with≡ pr1 | no with≡ pr2 | no with≡ pr3 rewrite pr1 | pr2 | pr3 = refl
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) | (r :: rs) with≡ pr1 | no with≡ pr2 | yes x with≡ pr3 rewrite pr1 | pr2 | pr3 | equalityCommutative pr1 = exFalso (noNotYes (transitivity (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (goPreservesCanonicalLeftOne a b)) (applyEquality (mapMaybe canonical) pr3)))
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) | (r :: rs) with≡ pr1 | yes m with≡ pr2 with inspect (go one (r :: rs) b)
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) | (r :: rs) with≡ pr1 | yes m with≡ pr2 | no with≡ x rewrite pr1 | pr2 | x | equalityCommutative pr1 = exFalso (noNotYes (transitivity (equalityCommutative (transitivity (goPreservesCanonicalLeftOne a b) (applyEquality (mapMaybe canonical) x))) (applyEquality (mapMaybe canonical) pr2)))
|
||||
goPreservesCanonicalLeftZero (zero :: a) (one :: b) | (r :: rs) with≡ pr1 | yes m with≡ pr2 | yes x₁ with≡ x rewrite pr1 | pr2 | x | equalityCommutative pr1 = applyEquality (λ i → yes (one :: i)) (yesInjective (transitivity (transitivity (applyEquality (mapMaybe canonical) (equalityCommutative pr2)) (goPreservesCanonicalLeftOne a b)) (applyEquality (mapMaybe canonical) x)))
|
||||
goPreservesCanonicalLeftZero (one :: a) [] rewrite equalityCommutative (canonicalIdempotent a) = refl
|
||||
goPreservesCanonicalLeftZero (one :: a) (zero :: b) with inspect (canonical a)
|
||||
goPreservesCanonicalLeftZero (one :: a) (zero :: b) | [] with≡ pr with inspect (go zero a b)
|
||||
goPreservesCanonicalLeftZero (one :: a) (zero :: b) | [] with≡ pr | no with≡ pr2 with inspect (go zero [] b)
|
||||
goPreservesCanonicalLeftZero (one :: a) (zero :: b) | [] with≡ pr | no with≡ pr2 | no with≡ pr3 rewrite pr | pr2 | pr3 = refl
|
||||
goPreservesCanonicalLeftZero (one :: a) (zero :: b) | [] with≡ pr | no with≡ pr2 | yes x with≡ pr3 rewrite pr | pr2 | pr3 = exFalso (noNotYes (transitivity (equalityCommutative (transitivity (transitivity (applyEquality (λ i → mapMaybe canonical (go zero i b)) (equalityCommutative pr)) (equalityCommutative (goPreservesCanonicalLeftZero a b))) (applyEquality (mapMaybe canonical) pr2))) (applyEquality (mapMaybe canonical) pr3)))
|
||||
goPreservesCanonicalLeftZero (one :: a) (zero :: b) | [] with≡ pr | yes x with≡ pr2 with inspect (go zero [] b)
|
||||
... | no with≡ pr3 rewrite pr | pr2 | pr3 = exFalso (noNotYes (transitivity (applyEquality (mapMaybe canonical) (equalityCommutative pr3)) (transitivity (applyEquality (λ i → mapMaybe canonical (go zero i b)) (equalityCommutative pr)) (transitivity (equalityCommutative (goPreservesCanonicalLeftZero a b)) (applyEquality (mapMaybe canonical) pr2)))))
|
||||
... | yes y with≡ pr3 rewrite pr | pr2 | pr3 = applyEquality (λ i → yes (one :: i)) (yesInjective (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (transitivity (transitivity (goPreservesCanonicalLeftZero a b) (applyEquality (λ i → mapMaybe canonical (go zero i b)) pr)) (applyEquality (mapMaybe canonical) pr3))))
|
||||
goPreservesCanonicalLeftZero (one :: a) (zero :: b) | (x :: y) with≡ pr with inspect (go zero a b)
|
||||
goPreservesCanonicalLeftZero (one :: a) (zero :: b) | (x :: y) with≡ pr | no with≡ pr2 with inspect (go zero (x :: y) b)
|
||||
... | no with≡ pr3 rewrite pr | pr2 | pr3 = refl
|
||||
... | yes z with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (transitivity (goPreservesCanonicalLeftZero a b) (applyEquality (mapMaybe canonical) pr3))))
|
||||
goPreservesCanonicalLeftZero (one :: a) (zero :: b) | (x :: y) with≡ pr | yes x₁ with≡ pr2 with inspect (go zero (x :: y) b)
|
||||
... | no with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = exFalso (noNotYes (transitivity (equalityCommutative (transitivity (goPreservesCanonicalLeftZero a b) (applyEquality (mapMaybe canonical) pr3))) (applyEquality (mapMaybe canonical) pr2)))
|
||||
... | yes z with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = applyEquality (λ i → yes (one :: i)) (yesInjective (transitivity (equalityCommutative (transitivity (equalityCommutative (goPreservesCanonicalLeftZero a b)) (applyEquality (mapMaybe canonical) pr2))) (applyEquality (mapMaybe canonical) pr3)))
|
||||
goPreservesCanonicalLeftZero (one :: a) (one :: b) with inspect (canonical a)
|
||||
goPreservesCanonicalLeftZero (one :: a) (one :: b) | [] with≡ x with inspect (go zero a b)
|
||||
goPreservesCanonicalLeftZero (one :: a) (one :: b) | [] with≡ pr | no with≡ pr2 with inspect (go zero [] b)
|
||||
... | no with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = refl
|
||||
... | yes z with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = exFalso (noNotYes (transitivity (applyEquality (mapMaybe canonical) (equalityCommutative pr2)) (transitivity (goPreservesCanonicalLeftZero a b) (applyEquality (mapMaybe canonical) pr3))))
|
||||
goPreservesCanonicalLeftZero (one :: a) (one :: b) | [] with≡ pr | yes y with≡ pr2 with inspect (go zero [] b)
|
||||
... | no with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = exFalso (noNotYes (transitivity (applyEquality (mapMaybe canonical) (equalityCommutative pr3)) (transitivity (equalityCommutative (goPreservesCanonicalLeftZero a b)) (applyEquality (mapMaybe canonical) pr2))))
|
||||
... | yes z with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = applyEquality yes v
|
||||
where
|
||||
t : canonical y ≡ canonical z
|
||||
t = yesInjective (transitivity (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (goPreservesCanonicalLeftZero a b)) (applyEquality (mapMaybe canonical) pr3))
|
||||
v : canonical (zero :: y) ≡ canonical (zero :: z)
|
||||
v with inspect (canonical y)
|
||||
v | [] with≡ pr4 rewrite pr4 | (transitivity (equalityCommutative t) pr4) = refl
|
||||
v | (x :: r) with≡ pr4 rewrite pr4 | transitivity (equalityCommutative t) pr4 = refl
|
||||
goPreservesCanonicalLeftZero (one :: a) (one :: b) | (r :: rs) with≡ x with inspect (go zero a b)
|
||||
goPreservesCanonicalLeftZero (one :: a) (one :: b) | (r :: rs) with≡ pr | no with≡ pr2 with inspect (go zero (r :: rs) b)
|
||||
goPreservesCanonicalLeftZero (one :: a) (one :: b) | (r :: rs) with≡ pr | no with≡ pr2 | no with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = refl
|
||||
goPreservesCanonicalLeftZero (one :: a) (one :: b) | (r :: rs) with≡ pr | no with≡ pr2 | yes z with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (transitivity (goPreservesCanonicalLeftZero a b) (applyEquality (mapMaybe canonical) pr3))))
|
||||
goPreservesCanonicalLeftZero (one :: a) (one :: b) | (r :: rs) with≡ pr | yes y with≡ pr2 with inspect (go zero (r :: rs) b)
|
||||
... | no with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr3)) (transitivity (equalityCommutative (goPreservesCanonicalLeftZero a b)) (applyEquality (mapMaybe canonical) pr2))))
|
||||
... | yes z with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = applyEquality yes v
|
||||
where
|
||||
t : canonical y ≡ canonical z
|
||||
t = yesInjective (transitivity (equalityCommutative (transitivity (equalityCommutative (goPreservesCanonicalLeftZero a b)) (applyEquality (mapMaybe canonical) pr2))) (applyEquality (mapMaybe canonical) pr3))
|
||||
v : canonical (zero :: y) ≡ canonical (zero :: z)
|
||||
v with inspect (canonical y)
|
||||
... | [] with≡ pr4 rewrite pr4 | transitivity (equalityCommutative t) pr4 = refl
|
||||
... | (r :: rs) with≡ pr4 rewrite pr4 | transitivity (equalityCommutative t) pr4 = refl
|
||||
|
||||
badLemma : (a b : BinNat) {t : BinNat} → go one a b ≡ yes t → canonical a ≡ [] → False
|
||||
badLemma a b pr1 pr2 with applyEquality (mapMaybe canonical) pr1
|
||||
badLemma a b pr1 pr2 | t with goPreservesCanonicalLeftOne a b
|
||||
... | th rewrite pr2 | t | goOneEmpty' b = noNotYes (equalityCommutative th)
|
||||
|
||||
goPreservesCanonicalLeftOne [] [] = refl
|
||||
goPreservesCanonicalLeftOne [] (x :: b) = refl
|
||||
goPreservesCanonicalLeftOne (zero :: a) [] with inspect (canonical a)
|
||||
goPreservesCanonicalLeftOne (zero :: a) [] | [] with≡ pr with inspect (go one a [])
|
||||
goPreservesCanonicalLeftOne (zero :: a) [] | [] with≡ pr | no with≡ x rewrite pr | x = refl
|
||||
goPreservesCanonicalLeftOne (zero :: a) [] | [] with≡ pr | yes y with≡ x rewrite pr | x = exFalso (noNotYes (transitivity (equalityCommutative (transitivity (goPreservesCanonicalLeftOne a []) (transitivity (applyEquality (λ i → mapMaybe canonical (go one i [])) pr) refl))) (applyEquality (mapMaybe canonical) x)))
|
||||
goPreservesCanonicalLeftOne (zero :: a) [] | (x :: xs) with≡ pr with inspect (go one a [])
|
||||
goPreservesCanonicalLeftOne (zero :: a) [] | (x :: xs) with≡ pr | no with≡ pr2 with inspect (go one (canonical a) [])
|
||||
... | no with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = refl
|
||||
... | yes z with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (transitivity (goPreservesCanonicalLeftOne a []) (applyEquality (mapMaybe canonical) pr3))))
|
||||
goPreservesCanonicalLeftOne (zero :: a) [] | (x :: xs) with≡ pr | yes y with≡ pr2 with inspect (go one (x :: xs) [])
|
||||
... | no with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr3)) (transitivity (equalityCommutative (goPreservesCanonicalLeftOne a [])) (applyEquality (mapMaybe canonical) pr2))))
|
||||
... | yes z with≡ pr3 rewrite pr | pr2 | pr3 | equalityCommutative pr = applyEquality (λ i → yes (one :: i)) (yesInjective (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (transitivity (goPreservesCanonicalLeftOne a []) (applyEquality (mapMaybe canonical) pr3))))
|
||||
goPreservesCanonicalLeftOne (one :: a) [] with inspect (canonical a)
|
||||
... | [] with≡ pr rewrite pr = refl
|
||||
... | (x :: xs) with≡ pr rewrite pr | equalityCommutative pr = applyEquality yes (transitivity (canonicalAscends' {zero} a λ pr2 → nonEmptyNotEmpty (transitivity (equalityCommutative pr) pr2)) (lemma a))
|
||||
where
|
||||
lemma : (a : BinNat) → canonical (zero :: a) ≡ canonical (zero :: canonical a)
|
||||
lemma [] = refl
|
||||
lemma (zero :: a) with inspect (canonical a)
|
||||
lemma (zero :: a) | [] with≡ pr rewrite pr = refl
|
||||
lemma (zero :: a) | (x :: bl) with≡ pr rewrite pr | equalityCommutative pr | equalityCommutative (canonicalIdempotent a) | pr = refl
|
||||
lemma (one :: a) rewrite equalityCommutative (canonicalIdempotent a) = refl
|
||||
goPreservesCanonicalLeftOne (zero :: as) (zero :: bs) with inspect (go one as bs)
|
||||
goPreservesCanonicalLeftOne (zero :: as) (zero :: bs) | no with≡ pr1 with inspect (canonical as)
|
||||
goPreservesCanonicalLeftOne (zero :: as) (zero :: bs) | no with≡ pr1 | [] with≡ pr2 rewrite pr1 | pr2 = refl
|
||||
goPreservesCanonicalLeftOne (zero :: as) (zero :: bs) | no with≡ pr1 | (a1 :: a) with≡ pr2 with inspect (go one (a1 :: a) bs)
|
||||
... | no with≡ pr3 rewrite pr1 | pr2 | pr3 = refl
|
||||
... | yes z with≡ pr3 rewrite pr1 | pr2 | pr3 = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr1)) (transitivity (transitivity (goPreservesCanonicalLeftOne as bs) (applyEquality (λ i → mapMaybe canonical (go one i bs)) pr2)) (applyEquality (mapMaybe canonical) pr3))))
|
||||
goPreservesCanonicalLeftOne (zero :: as) (zero :: bs) | yes x with≡ pr1 with inspect (canonical as)
|
||||
goPreservesCanonicalLeftOne (zero :: as) (zero :: bs) | yes x with≡ pr1 | [] with≡ pr2 rewrite pr1 | pr2 = exFalso (badLemma as bs pr1 pr2)
|
||||
goPreservesCanonicalLeftOne (zero :: as) (zero :: bs) | yes x with≡ pr1 | (r :: rs) with≡ pr2 with inspect (go one (r :: rs) bs)
|
||||
... | no with≡ pr3 rewrite pr1 | pr2 | pr3 = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr3)) (transitivity (equalityCommutative (transitivity (goPreservesCanonicalLeftOne as bs) (applyEquality (λ i → mapMaybe canonical (go one i bs)) pr2))) (applyEquality (mapMaybe canonical) pr1))))
|
||||
... | yes z with≡ pr3 rewrite pr1 | pr2 | pr3 = applyEquality (λ i → yes (one :: i)) (yesInjective (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr1)) (transitivity (transitivity (goPreservesCanonicalLeftOne as bs) (applyEquality (λ i → mapMaybe canonical (go one i bs)) pr2)) (applyEquality (mapMaybe canonical) pr3))))
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) with inspect (go one as bs)
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) | no with≡ pr with inspect (canonical as)
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) | no with≡ pr | [] with≡ pr2 rewrite pr | pr2 = refl
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) | no with≡ pr | (r :: rs) with≡ pr2 with inspect (go one (r :: rs) bs)
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) | no with≡ pr | (r :: rs) with≡ pr2 | no with≡ pr3 rewrite pr | pr2 | pr3 = refl
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) | no with≡ pr | (r :: rs) with≡ pr2 | yes z with≡ pr3 rewrite pr | pr2 | pr3 = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr)) (transitivity (transitivity (goPreservesCanonicalLeftOne as bs) (applyEquality (λ i → mapMaybe canonical (go one i bs)) pr2)) (applyEquality (mapMaybe canonical) pr3))))
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) | yes x₁ with≡ pr with inspect (canonical as)
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) | yes x₁ with≡ pr | [] with≡ pr2 rewrite pr | pr2 = exFalso (badLemma as bs pr pr2)
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) | yes x₁ with≡ pr | (r :: rs) with≡ pr2 with inspect (go one (r :: rs) bs)
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) | yes x₁ with≡ pr | (r :: rs) with≡ pr2 | no with≡ pr3 rewrite pr | pr2 | pr3 = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr3)) (transitivity (transitivity (applyEquality (λ i → mapMaybe canonical (go one i bs)) (equalityCommutative pr2)) (equalityCommutative (goPreservesCanonicalLeftOne as bs))) (applyEquality (mapMaybe canonical) pr))))
|
||||
goPreservesCanonicalLeftOne (zero :: as) (one :: bs) | yes x with≡ pr | (r :: rs) with≡ pr2 | yes y with≡ pr3 rewrite pr | pr2 | pr3 = applyEquality yes v
|
||||
where
|
||||
t : canonical x ≡ canonical y
|
||||
t = yesInjective (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr)) (transitivity (transitivity (goPreservesCanonicalLeftOne as bs) (applyEquality (λ i → mapMaybe canonical (go one i bs)) pr2)) (applyEquality (mapMaybe canonical) pr3)))
|
||||
v : canonical (zero :: x) ≡ canonical (zero :: y)
|
||||
v with inspect (canonical x)
|
||||
... | [] with≡ pr3 rewrite pr3 | transitivity (equalityCommutative t) pr3 = refl
|
||||
... | (r :: rs) with≡ pr3 rewrite pr3 | transitivity (equalityCommutative t) pr3 = refl
|
||||
goPreservesCanonicalLeftOne (one :: as) (zero :: bs) with inspect (go zero as bs)
|
||||
goPreservesCanonicalLeftOne (one :: as) (zero :: bs) | no with≡ pr with inspect (go zero (canonical as) bs)
|
||||
goPreservesCanonicalLeftOne (one :: as) (zero :: bs) | no with≡ pr | no with≡ pr2 rewrite pr | pr2 = refl
|
||||
goPreservesCanonicalLeftOne (one :: as) (zero :: bs) | no with≡ pr | yes x with≡ pr2 rewrite pr | pr2 = exFalso (noNotYes (transitivity (applyEquality (mapMaybe canonical) (equalityCommutative pr)) (transitivity (goPreservesCanonicalLeftZero as bs) (applyEquality (mapMaybe canonical) pr2))))
|
||||
goPreservesCanonicalLeftOne (one :: as) (zero :: bs) | yes x with≡ pr with inspect (go zero (canonical as) bs)
|
||||
goPreservesCanonicalLeftOne (one :: as) (zero :: bs) | yes x with≡ pr | yes y with≡ pr2 rewrite pr | pr2 = applyEquality yes v
|
||||
where
|
||||
t : canonical x ≡ canonical y
|
||||
t = yesInjective (transitivity (equalityCommutative (transitivity (equalityCommutative (goPreservesCanonicalLeftZero as bs)) (applyEquality (mapMaybe canonical) pr))) (applyEquality (mapMaybe canonical) pr2))
|
||||
v : canonical (zero :: x) ≡ canonical (zero :: y)
|
||||
v with inspect (canonical x)
|
||||
... | [] with≡ pr3 rewrite pr3 | transitivity (equalityCommutative t) pr3 = refl
|
||||
... | (r :: rs) with≡ pr3 rewrite pr3 | transitivity (equalityCommutative t) pr3 = refl
|
||||
goPreservesCanonicalLeftOne (one :: as) (zero :: bs) | yes x with≡ pr | no with≡ pr2 rewrite pr | pr2 = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (transitivity (equalityCommutative (goPreservesCanonicalLeftZero as bs)) (applyEquality (mapMaybe canonical) pr))))
|
||||
goPreservesCanonicalLeftOne (one :: as) (one :: bs) with inspect (go one as bs)
|
||||
goPreservesCanonicalLeftOne (one :: as) (one :: bs) | y with≡ pr with inspect (go one (canonical as) bs)
|
||||
goPreservesCanonicalLeftOne (one :: as) (one :: bs) | no with≡ pr | no with≡ pr2 rewrite pr | pr2 = refl
|
||||
goPreservesCanonicalLeftOne (one :: as) (one :: bs) | no with≡ pr | yes x₁ with≡ pr2 rewrite pr | pr2 = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr)) (transitivity (goPreservesCanonicalLeftOne as bs) (applyEquality (mapMaybe canonical) pr2))))
|
||||
goPreservesCanonicalLeftOne (one :: as) (one :: bs) | yes x₁ with≡ pr | no with≡ pr2 rewrite pr | pr2 = exFalso (noNotYes (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr2)) (transitivity (equalityCommutative (goPreservesCanonicalLeftOne as bs)) (applyEquality (mapMaybe canonical) pr))))
|
||||
goPreservesCanonicalLeftOne (one :: as) (one :: bs) | yes x₁ with≡ pr | yes x₂ with≡ pr2 rewrite pr | pr2 = applyEquality (λ i → yes (one :: i)) (yesInjective (transitivity (transitivity (equalityCommutative (applyEquality (mapMaybe canonical) pr)) (goPreservesCanonicalLeftOne as bs)) (applyEquality (mapMaybe canonical) pr2)))
|
||||
|
||||
goPreservesCanonicalLeft : (state : Bit) → (a b : BinNat) → mapMaybe canonical (go state a b) ≡ mapMaybe canonical (go state (canonical a) b)
|
||||
goPreservesCanonicalLeft zero = goPreservesCanonicalLeftZero
|
||||
goPreservesCanonicalLeft one = goPreservesCanonicalLeftOne
|
@@ -2,7 +2,6 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
open import Maybe
|
||||
open import Numbers.BinaryNaturals.SubtractionGo
|
||||
|
@@ -4,7 +4,6 @@ open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Definition
|
||||
open import Semirings.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.Definition
|
||||
open import Setoids.Setoids
|
||||
|
@@ -2,13 +2,7 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Multiplication
|
||||
open import Numbers.Integers.Definition
|
||||
open import Numbers.Integers.Addition
|
||||
open import Semirings.Definition
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Setoids.Setoids
|
||||
|
||||
module Numbers.Integers.Multiplication where
|
||||
|
||||
|
@@ -3,15 +3,10 @@
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Integers.Definition
|
||||
open import Numbers.Integers.RingStructure.Ring
|
||||
open import Numbers.Integers.Addition
|
||||
open import Numbers.Integers.Multiplication
|
||||
open import Semirings.Definition
|
||||
open import Rings.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Orders.Total.Definition
|
||||
open import Orders.Partial.Definition
|
||||
|
@@ -9,9 +9,6 @@ open import Numbers.Integers.Order
|
||||
open import Numbers.Integers.RingStructure.Ring
|
||||
open import Numbers.Integers.RingStructure.IntegralDomain
|
||||
open import Semirings.Definition
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Rings.EuclideanDomains.Definition
|
||||
open import Orders.Total.Definition
|
||||
open import Numbers.Naturals.EuclideanAlgorithm
|
||||
|
@@ -2,12 +2,7 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Multiplication
|
||||
open import Numbers.Integers.RingStructure.Ring
|
||||
open import Semirings.Definition
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Rings.IntegralDomains.Definition
|
||||
|
||||
module Numbers.Integers.RingStructure.IntegralDomain where
|
||||
|
@@ -4,7 +4,6 @@ open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Multiplication
|
||||
open import Semirings.Definition
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Setoids.Setoids
|
||||
|
||||
|
@@ -3,16 +3,8 @@
|
||||
-- a natural together with a proof that it is small.
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet.Definition
|
||||
open import Functions
|
||||
open import Semirings.Definition
|
||||
open import Orders.Total.Definition
|
||||
open import Numbers.Modulo.Definition
|
||||
|
@@ -1,19 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet.Definition
|
||||
open import Functions
|
||||
open import Semirings.Definition
|
||||
open import Numbers.Naturals.EuclideanAlgorithm
|
||||
open import Orders.Total.Definition
|
||||
|
||||
module Numbers.Modulo.Definition where
|
||||
|
||||
|
@@ -2,23 +2,19 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet.Definition
|
||||
open import Sets.FinSet.Lemmas
|
||||
open import Sets.Cardinality.Finite.Definition
|
||||
open import Functions
|
||||
open import Semirings.Definition
|
||||
open import Numbers.Modulo.Definition
|
||||
open import Numbers.Modulo.Addition
|
||||
open import Orders.Total.Definition
|
||||
open import Numbers.Naturals.EuclideanAlgorithm
|
||||
open import Numbers.Modulo.ModuloFunction
|
||||
|
||||
module Numbers.Modulo.Group where
|
||||
|
@@ -7,10 +7,7 @@ open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Numbers.Naturals.Order.WellFounded
|
||||
open import Orders.WellFounded.Induction
|
||||
open import KeyValue.KeyValue
|
||||
open import Orders.Total.Definition
|
||||
open import Vectors
|
||||
open import Maybe
|
||||
open import Semirings.Definition
|
||||
|
||||
module Numbers.Naturals.EuclideanAlgorithm where
|
||||
|
@@ -1,17 +1,11 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Addition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Multiplication
|
||||
open import Numbers.Naturals.Exponentiation
|
||||
open import Numbers.Naturals.Subtraction
|
||||
open import Semirings.Definition
|
||||
open import Monoids.Definition
|
||||
open import Orders.Total.Definition
|
||||
|
||||
module Numbers.Naturals.Naturals where
|
||||
|
@@ -2,7 +2,6 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Semirings.Definition
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Orders.Total.Definition
|
||||
open import Orders.Partial.Definition
|
||||
|
@@ -1,13 +1,10 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Orders.WellFounded.Definition
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Semirings.Definition
|
||||
open import Orders.Total.Definition
|
||||
|
||||
module Numbers.Naturals.Order.WellFounded where
|
||||
open Semiring ℕSemiring
|
||||
|
@@ -1,8 +1,6 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Addition
|
||||
open import Numbers.Naturals.Multiplication
|
||||
|
Some files were not shown because too many files have changed in this diff Show More
Reference in New Issue
Block a user