# HG changeset patch # User Shinji KONO # Date 1704100896 -32400 # Node ID fa52d72f4bb3646b991d98f8d9bad8234702fa9c # Parent d1b6fb58aad00be04c0de41be80241b41daebe82 ... diff -r d1b6fb58aad0 -r fa52d72f4bb3 ZF.agda-lib --- a/ZF.agda-lib Sat Aug 26 10:36:09 2023 +0900 +++ b/ZF.agda-lib Mon Jan 01 18:21:36 2024 +0900 @@ -1,3 +1,6 @@ name: ZF depend: standard-library include: src +flags: + --warning=noUnsupportedIndexedMatch + diff -r d1b6fb58aad0 -r fa52d72f4bb3 src/OD.agda --- a/src/OD.agda Sat Aug 26 10:36:09 2023 +0900 +++ b/src/OD.agda Mon Jan 01 18:21:36 2024 +0900 @@ -1,7 +1,8 @@ -{-# OPTIONS --allow-unsolved-metas #-} +{-# OPTIONS --cubical-compatible --safe #-} open import Level open import Ordinals -module OD {n : Level } (O : Ordinals {n} ) where +import HODBase +module OD {n : Level } (O : Ordinals {n} ) (HODAxiom : HODBase.ODAxiom O) where open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) open import Relation.Binary.PropositionalEquality hiding ( [_] ) @@ -23,87 +24,25 @@ -- Ordinal Definable Set -record OD : Set (suc n ) where - field - def : (x : Ordinal ) → Set n - -open OD +open HODBase.HOD +open HODBase.OD open _∧_ open _∨_ open Bool -record _==_ ( a b : OD ) : Set n where - field - eq→ : ∀ { x : Ordinal } → def a x → def b x - eq← : ∀ { x : Ordinal } → def b x → def a x - -==-refl : { x : OD } → x == x -==-refl {x} = record { eq→ = λ x → x ; eq← = λ x → x } - -open _==_ +open HODBase._==_ -==-trans : { x y z : OD } → x == y → y == z → x == z -==-trans x=y y=z = record { eq→ = λ {m} t → eq→ y=z (eq→ x=y t) ; eq← = λ {m} t → eq← x=y (eq← y=z t) } - -==-sym : { x y : OD } → x == y → y == x -==-sym x=y = record { eq→ = λ {m} t → eq← x=y t ; eq← = λ {m} t → eq→ x=y t } - - -⇔→== : { x y : OD } → ( {z : Ordinal } → (def x z ⇔ def y z)) → x == y -eq→ ( ⇔→== {x} {y} eq ) {z} m = proj1 eq m -eq← ( ⇔→== {x} {y} eq ) {z} m = proj2 eq m +open HODBase.ODAxiom HODAxiom --- --- OD is an equation on Ordinals, so it contains Ordinals. If these Ordinals have one-to-one --- correspondence to the OD then the OD looks like a ZF Set. --- --- If all ZF Set have supreme upper bound, the solutions of OD have to be bounded, i.e. --- bbounded ODs are ZF Set. Unbounded ODs are classes. --- --- In classical Set Theory, HOD is used, as a subset of OD, --- HOD = { x | TC x ⊆ OD } --- where TC x is a transitive clusure of x, i.e. Union of all elemnts of all subset of x. --- This is not possible because we don't have V yet. So we assumes HODs are bounded OD. --- --- We also assumes HODs are isomorphic to Ordinals, which is ususally proved by Goedel number tricks. --- There two contraints on the HOD order, one is ∋, the other one is ⊂. --- ODs have an ovbious maximum, but Ordinals are not, but HOD has no maximum because there is an aribtrary --- bound on each HOD. --- --- In classical Set Theory, sup is defined by Uion, since we are working on constructive logic, --- we need explict assumption on sup for unrestricted Replacement. --- --- ==→o≡ is necessary to prove axiom of extensionality. - --- Ordinals in OD , the maximum -Ords : OD -Ords = record { def = λ x → Lift n ⊤ } - -record HOD : Set (suc n) where - field - od : OD - odmax : Ordinal - ¬a ¬b c = ⊥-elim ( o<> (⊆→o≤ {x , x} {y} y⊆x,x ) lemma1 ) where lemma : {z : Ordinal} → (z ≡ & x) ∨ (z ≡ & x) → & x ≡ z @@ -298,20 +234,25 @@ ε-induction : { ψ : HOD → Set (suc n)} → ( {x : HOD } → ({ y : HOD } → x ∋ y → ψ y ) → ψ x ) → (x : HOD ) → ψ x -ε-induction {ψ} ind x = subst (λ k → ψ k ) *iso (ε-induction-ord (osuc (& x)) <-osuc ) where - induction : (ox : Ordinal) → ((oy : Ordinal) → oy o< ox → ψ (* oy)) → ψ (* ox) - induction ox prev = ind ( λ {y} lt → subst (λ k → ψ k ) *iso (prev (& y) (o<-subst (c<→o< lt) refl &iso ))) - ε-induction-ord : (ox : Ordinal) { oy : Ordinal } → oy o< ox → ψ (* oy) - ε-induction-ord ox {oy} lt = TransFinite {λ oy → ψ (* oy)} induction oy +ε-induction {ψ} ind x = ε-induction-hod _ {& x} <-osuc x <-osuc where + induction2 : (x₁ : Ordinal) → + ((y : Ordinal) → y o< x₁ → (y₁ : HOD) → & y₁ o< osuc y → ψ y₁) → + (y : HOD) → & y o< osuc x₁ → ψ y + induction2 x prev y y≤x = ind (λ {y₁} lt → prev (& y₁) (lemma1 y₁ lt) y₁ <-osuc ) where + lemma1 : (y₁ : HOD) → y ∋ y₁ → & y₁ o< x + lemma1 y₁ lt with trio< (& y₁) x + ... | tri< a ¬b ¬c = a + ... | tri> ¬a ¬b c = ⊥-elim (o≤> (ordtrans (c<→o< lt) y≤x) c ) + ... | tri≈ ¬a b ¬c with osuc-≡< y≤x + ... | case1 y=x = subst (λ k → & y₁ o< k ) y=x (c<→o< lt) + ... | case2 y ¬a ¬b c = osuc x -omax x y | tri≈ ¬a refl ¬c = osuc x +omax x y | tri≈ ¬a b ¬c = osuc x omax< : ( x y : Ordinal ) → x o< y → osuc y ≡ omax x y omax< x y lt with trio< x y @@ -190,7 +192,7 @@ omxx x with trio< x x omxx x | tri< a ¬b ¬c = ⊥-elim (¬b refl ) omxx x | tri> ¬a ¬b c = ⊥-elim (¬b refl ) -omxx x | tri≈ ¬a refl ¬c = refl +omxx x | tri≈ ¬a b ¬c = refl omxxx : ( x : Ordinal ) → omax x (omax x x ) ≡ osuc (osuc x) omxxx x = trans ( cong ( λ k → omax x k ) (omxx x )) (sym ( omax< x (osuc x) <-osuc )) diff -r d1b6fb58aad0 -r fa52d72f4bb3 src/Ordinals.agda --- a/src/Ordinals.agda Sat Aug 26 10:36:09 2023 +0900 +++ b/src/Ordinals.agda Mon Jan 01 18:21:36 2024 +0900 @@ -1,3 +1,5 @@ +{-# OPTIONS --cubical-compatible --safe #-} + open import Level module Ordinals where diff -r d1b6fb58aad0 -r fa52d72f4bb3 src/Tychonoff.agda --- a/src/Tychonoff.agda Sat Aug 26 10:36:09 2023 +0900 +++ b/src/Tychonoff.agda Mon Jan 01 18:21:36 2024 +0900 @@ -376,7 +376,7 @@ import Axiom.Extensionality.Propositional postulate f-extensionality : { n m : Level} → Axiom.Extensionality.Propositional.Extensionality n m -open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) +-- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) -- -- We have UFLP both in P and Q. Given an ultra filter F on P x Q. It has limits on P and Q because a projection of ultra filter diff -r d1b6fb58aad0 -r fa52d72f4bb3 src/ZProduct.agda --- a/src/ZProduct.agda Sat Aug 26 10:36:09 2023 +0900 +++ b/src/ZProduct.agda Mon Jan 01 18:21:36 2024 +0900 @@ -340,7 +340,7 @@ -- Set of All function from A to B -- -open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) +-- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) Funcs : (A B : HOD) → HOD Funcs A B = record { od = record { def = λ x → FuncHOD A B x } ; odmax = osuc (& (ZFP A B)) @@ -361,6 +361,7 @@ record Injection (A B : Ordinal ) : Set n where field i→ : (x : Ordinal ) → odef (* A) x → Ordinal + irr : (x : Ordinal ) → ( lt1 lt2 : odef (* A) x ) → i→ x lt1 ≡ i→ x lt2 iB : (x : Ordinal ) → ( lt : odef (* A) x ) → odef (* B) ( i→ x lt ) inject : (x y : Ordinal ) → ( ltx : odef (* A) x ) ( lty : odef (* A) y ) → i→ x ltx ≡ i→ y lty → x ≡ y diff -r d1b6fb58aad0 -r fa52d72f4bb3 src/bijection.agda --- a/src/bijection.agda Sat Aug 26 10:36:09 2023 +0900 +++ b/src/bijection.agda Mon Jan 01 18:21:36 2024 +0900 @@ -47,6 +47,27 @@ bi-inject→ : {n m : Level} {R : Set n} {S : Set m} → (rs : Bijection R S) → {x y : R} → fun→ rs x ≡ fun→ rs y → x ≡ y bi-inject→ rs {x} {y} eq = subst₂ (λ j k → j ≡ k ) (fiso← rs _) (fiso← rs _) (cong (fun← rs) eq) +bi-∨ : {n m n1 m1 : Level } {A : Set n} {B : Set m} {C : Set n1} {D : Set m1} → (ab : Bijection A B) → (cd : Bijection C D ) + → Bijection (A ∨ C) (B ∨ D) +bi-∨ {_} {_} {_} {_} {A} {B} {C} {D} ab cd = record { + fun→ = fun→1 + ; fun← = fun←1 + ; fiso→ = fiso→1 + ; fiso← = fiso←1 + } where + fun→1 : (A ∨ C) → (B ∨ D) + fun→1 (case1 a) = case1 (fun→ ab a) + fun→1 (case2 c) = case2 (fun→ cd c) + fun←1 : (B ∨ D) → (A ∨ C) + fun←1 (case1 a) = case1 (fun← ab a) + fun←1 (case2 c) = case2 (fun← cd c) + fiso→1 : (x : B ∨ D) → fun→1 (fun←1 x) ≡ x + fiso→1 (case1 a) = cong case1 (fiso→ ab a) + fiso→1 (case2 c) = cong case2 (fiso→ cd c) + fiso←1 : (x : A ∨ C) → fun←1 (fun→1 x) ≡ x + fiso←1 (case1 a) = cong case1 (fiso← ab a) + fiso←1 (case2 c) = cong case2 (fiso← cd c) + open import Relation.Binary.Structures bijIsEquivalence : {n : Level } → IsEquivalence (Bijection {n} {n}) @@ -511,6 +532,42 @@ a : A fa=c : f a ≡ c +record IsImage0 (A B : Set ) (f : (x : A ) → B) (x : B ) : Set where + field + y : A + x=fy : x ≡ f y + +IsImage : (a b : Set) (iab : InjectiveF a b ) (x : b ) → Set +IsImage a b iab x = IsImage0 a b (InjectiveF.f iab) x + +Bernstein : (A B : Set) + → (fi : InjectiveF A B ) → (gi : InjectiveF B A ) + → (is-A : (b : B ) → Dec (Is A B (InjectiveF.f fi) b) ) + → (is-B : (a : A ) → Dec (Is B A (InjectiveF.f gi) a) ) + → Bijection A B +Bernstein A B fi gi isa isb = ? where + open InjectiveF + gfi : InjectiveF A A + gfi = record { f = λ x → f gi (f fi x) ; inject = λ {x} {y} eq → inject fi (inject gi eq) } + data gfImage : (x : A) → Set where + a-g : {x : A} → (¬ib : ¬ ( IsImage B A gi x )) → gfImage x + next-gf : {x : A} → (ix : IsImage A A gfi x) → (gfiy : gfImage (IsImage0.y ix) ) → gfImage x + data ¬gfImage : (x : A) → Set where + ngf : {x : A} → (¬gfiy : ¬ gfImage x) → ¬gfImage x + gf02 : {x : A} → IsImage B A gi x ∨ (¬ IsImage B A gi x) ∨ ((ix : IsImage A A gfi x) → ¬ gfImage (IsImage0.y ix) ) → ¬ gfImage x + gf02 {x} c gf = ? + gfi∨ : (x : A) → gfImage x ∨ ¬gfImage x + gfi∨ x with isb x + ... | no ¬ib = case1 ( a-g (λ ib → ¬ib (record { a = IsImage0.y ib ; fa=c = sym (IsImage0.x=fy ib) }))) + ... | yes ib with isa (f fi x) + ... | no ¬ia = case2 ( ngf ? ) + ... | yes ia = case1 ( next-gf record { y = f gi (Is.a ib) ; x=fy = br00 } (a-g br01 ) ) where + br00 : x ≡ f gi (f fi (f gi (Is.a ib))) + br00 = ? + br01 : ¬ IsImage B A gi (f gi (Is.a ib)) + br01 record { y = y ; x=fy = x=fy } = ? + + Countable-Bernstein : (A B C : Set) → Bijection A ℕ → Bijection C ℕ → (fi : InjectiveF A B ) → (gi : InjectiveF B C ) → (is-A : (c : C ) → Dec (Is A C (λ x → (InjectiveF.f gi (InjectiveF.f fi x))) c )) diff -r d1b6fb58aad0 -r fa52d72f4bb3 src/logic.agda --- a/src/logic.agda Sat Aug 26 10:36:09 2023 +0900 +++ b/src/logic.agda Mon Jan 01 18:21:36 2024 +0900 @@ -1,3 +1,5 @@ +{-# OPTIONS --cubical-compatible --safe #-} + module logic where open import Level @@ -5,14 +7,11 @@ open import Relation.Binary hiding (_⇔_ ) open import Data.Empty + data Bool : Set where true : Bool false : Bool -data Two : Set where - i0 : Two - i1 : Two - record _∧_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where constructor ⟪_,_⟫ field @@ -26,13 +25,6 @@ _⇔_ : {n m : Level } → ( A : Set n ) ( B : Set m ) → Set (n ⊔ m) _⇔_ A B = ( A → B ) ∧ ( B → A ) -∧-exch : {n m : Level} {A : Set n} { B : Set m } → A ∧ B → B ∧ A -∧-exch p = ⟪ _∧_.proj2 p , _∧_.proj1 p ⟫ - -∨-exch : {n m : Level} {A : Set n} { B : Set m } → A ∨ B → B ∨ A -∨-exch (case1 x) = case2 x -∨-exch (case2 x) = case1 x - contra-position : {n m : Level } {A : Set n} {B : Set m} → (A → B) → ¬ B → ¬ A contra-position {n} {m} {A} {B} f ¬b a = ¬b ( f a ) @@ -46,10 +38,6 @@ de-morgan {n} {A} {B} and (case1 ¬A) = ⊥-elim ( ¬A ( _∧_.proj1 and )) de-morgan {n} {A} {B} and (case2 ¬B) = ⊥-elim ( ¬B ( _∧_.proj2 and )) -de-morgan∨ : {n : Level } {A B : Set n} → A ∨ B → ¬ ( (¬ A ) ∧ (¬ B ) ) -de-morgan∨ {n} {A} {B} (case1 a) and = ⊥-elim ( _∧_.proj1 and a ) -de-morgan∨ {n} {A} {B} (case2 b) and = ⊥-elim ( _∧_.proj2 and b ) - dont-or : {n m : Level} {A : Set n} { B : Set m } → A ∨ B → ¬ A → B dont-or {A} {B} (case1 a) ¬A = ⊥-elim ( ¬A a ) dont-or {A} {B} (case2 b) ¬A = b @@ -79,11 +67,10 @@ false <=> false = true _ <=> _ = false -open import Relation.Binary.PropositionalEquality +infixr 130 _\/_ +infixr 140 _/\_ -not-not-bool : { b : Bool } → not (not b) ≡ b -not-not-bool {true} = refl -not-not-bool {false} = refl +open import Relation.Binary.PropositionalEquality record Bijection {n m : Level} (R : Set n) (S : Set m) : Set (n Level.⊔ m) where field @@ -96,13 +83,14 @@ injection R S f = (x y : R) → f x ≡ f y → x ≡ y +not-not-bool : { b : Bool } → not (not b) ≡ b +not-not-bool {true} = refl +not-not-bool {false} = refl + ¬t=f : (t : Bool ) → ¬ ( not t ≡ t) ¬t=f true () ¬t=f false () -infixr 130 _\/_ -infixr 140 _/\_ - ≡-Bool-func : {A B : Bool } → ( A ≡ true → B ≡ true ) → ( B ≡ true → A ≡ true ) → A ≡ B ≡-Bool-func {true} {true} a→b b→a = refl ≡-Bool-func {false} {true} a→b b→a with b→a refl @@ -129,89 +117,57 @@ ¬-bool refl () lemma-∧-0 : {a b : Bool} → a /\ b ≡ true → a ≡ false → ⊥ -lemma-∧-0 {true} {true} refl () lemma-∧-0 {true} {false} () lemma-∧-0 {false} {true} () lemma-∧-0 {false} {false} () +lemma-∧-0 {true} {true} eq1 () lemma-∧-1 : {a b : Bool} → a /\ b ≡ true → b ≡ false → ⊥ -lemma-∧-1 {true} {true} refl () lemma-∧-1 {true} {false} () lemma-∧-1 {false} {true} () lemma-∧-1 {false} {false} () +lemma-∧-1 {true} {true} eq1 () bool-and-tt : {a b : Bool} → a ≡ true → b ≡ true → ( a /\ b ) ≡ true bool-and-tt refl refl = refl bool-∧→tt-0 : {a b : Bool} → ( a /\ b ) ≡ true → a ≡ true -bool-∧→tt-0 {true} {true} refl = refl +bool-∧→tt-0 {true} {true} eq = refl bool-∧→tt-0 {false} {_} () bool-∧→tt-1 : {a b : Bool} → ( a /\ b ) ≡ true → b ≡ true -bool-∧→tt-1 {true} {true} refl = refl +bool-∧→tt-1 {true} {true} eq = refl bool-∧→tt-1 {true} {false} () bool-∧→tt-1 {false} {false} () bool-or-1 : {a b : Bool} → a ≡ false → ( a \/ b ) ≡ b -bool-or-1 {false} {true} refl = refl -bool-or-1 {false} {false} refl = refl +bool-or-1 {false} {true} eq = refl +bool-or-1 {false} {false} eq = refl bool-or-2 : {a b : Bool} → b ≡ false → (a \/ b ) ≡ a -bool-or-2 {true} {false} refl = refl -bool-or-2 {false} {false} refl = refl +bool-or-2 {true} {false} eq = refl +bool-or-2 {false} {false} eq = refl bool-or-3 : {a : Bool} → ( a \/ true ) ≡ true bool-or-3 {true} = refl bool-or-3 {false} = refl bool-or-31 : {a b : Bool} → b ≡ true → ( a \/ b ) ≡ true -bool-or-31 {true} {true} refl = refl -bool-or-31 {false} {true} refl = refl +bool-or-31 {true} {true} eq = refl +bool-or-31 {false} {true} eq = refl bool-or-4 : {a : Bool} → ( true \/ a ) ≡ true bool-or-4 {true} = refl bool-or-4 {false} = refl bool-or-41 : {a b : Bool} → a ≡ true → ( a \/ b ) ≡ true -bool-or-41 {true} {b} refl = refl +bool-or-41 {true} {b} eq = refl bool-and-1 : {a b : Bool} → a ≡ false → (a /\ b ) ≡ false -bool-and-1 {false} {b} refl = refl +bool-and-1 {false} {b} eq = refl bool-and-2 : {a b : Bool} → b ≡ false → (a /\ b ) ≡ false -bool-and-2 {true} {false} refl = refl -bool-and-2 {false} {false} refl = refl +bool-and-2 {true} {false} eq = refl +bool-and-2 {false} {false} eq = refl +bool-and-2 {true} {true} () +bool-and-2 {false} {true} () -open import Data.Nat -open import Data.Nat.Properties - -_≥b_ : ( x y : ℕ) → Bool -x ≥b y with <-cmp x y -... | tri< a ¬b ¬c = false -... | tri≈ ¬a b ¬c = true -... | tri> ¬a ¬b c = true - -_>b_ : ( x y : ℕ) → Bool -x >b y with <-cmp x y -... | tri< a ¬b ¬c = false -... | tri≈ ¬a b ¬c = false -... | tri> ¬a ¬b c = true - -_≤b_ : ( x y : ℕ) → Bool -x ≤b y = y ≥b x - -_b x - -open import Relation.Binary.PropositionalEquality - -¬i0≡i1 : ¬ ( i0 ≡ i1 ) -¬i0≡i1 () - -¬i0→i1 : {x : Two} → ¬ (x ≡ i0 ) → x ≡ i1 -¬i0→i1 {i0} ne = ⊥-elim ( ne refl ) -¬i0→i1 {i1} ne = refl - -¬i1→i0 : {x : Two} → ¬ (x ≡ i1 ) → x ≡ i0 -¬i1→i0 {i0} ne = refl -¬i1→i0 {i1} ne = ⊥-elim ( ne refl ) - diff -r d1b6fb58aad0 -r fa52d72f4bb3 src/nat.agda --- a/src/nat.agda Sat Aug 26 10:36:09 2023 +0900 +++ b/src/nat.agda Mon Jan 01 18:21:36 2024 +0900 @@ -1,4 +1,5 @@ -{-# OPTIONS --allow-unsolved-metas #-} +{-# OPTIONS --cubical-compatible --safe #-} + module nat where open import Data.Nat @@ -104,15 +105,15 @@ div2-eq : (x : ℕ ) → div2-rev ( div2 x ) ≡ x div2-eq zero = refl div2-eq (suc zero) = refl -div2-eq (suc (suc x)) with div2 x | inspect div2 x -... | ⟪ x1 , true ⟫ | record { eq = eq1 } = begin -- eq1 : div2 x ≡ ⟪ x1 , true ⟫ +div2-eq (suc (suc x)) with div2 x in eq1 +... | ⟪ x1 , true ⟫ = begin -- eq1 : div2 x ≡ ⟪ x1 , true ⟫ div2-rev ⟪ suc x1 , true ⟫ ≡⟨⟩ suc (suc (x1 + suc x1)) ≡⟨ cong (λ k → suc (suc k )) (+-comm x1 _ ) ⟩ suc (suc (suc (x1 + x1))) ≡⟨⟩ suc (suc (div2-rev ⟪ x1 , true ⟫)) ≡⟨ cong (λ k → suc (suc (div2-rev k ))) (sym eq1) ⟩ suc (suc (div2-rev (div2 x))) ≡⟨ cong (λ k → suc (suc k)) (div2-eq x) ⟩ suc (suc x) ∎ where open ≡-Reasoning -... | ⟪ x1 , false ⟫ | record { eq = eq1 } = begin +... | ⟪ x1 , false ⟫ = begin div2-rev ⟪ suc x1 , false ⟫ ≡⟨⟩ suc (x1 + suc x1) ≡⟨ cong (λ k → (suc k )) (+-comm x1 _ ) ⟩ suc (suc (x1 + x1)) ≡⟨⟩ @@ -138,6 +139,40 @@ _-_ = minus +sn-m=sn-m : {m n : ℕ } → m ≤ n → suc n - m ≡ suc ( n - m ) +sn-m=sn-m {0} {n} z≤n = refl +sn-m=sn-m {suc m} {suc n} (s≤s m ¬a ¬b c = case2 ( 0 → m > 0 → (d : Dividable k m ) → Dividable.factor d > 0 -00 m>0 d with Dividable.factor d | inspect Dividable.factor d -... | zero | record { eq = eq1 } = ⊥-elim ( nat-≡< ff1 m>0 ) where +00 m>0 d with Dividable.factor d in eq1 +... | zero = ⊥-elim ( nat-≡< ff1 m>0 ) where ff1 : 0 ≡ m ff1 = begin 0 ≡⟨⟩ 0 * k + 0 ≡⟨ cong (λ j → j * k + 0) (sym eq1) ⟩ Dividable.factor d * k + 0 ≡⟨ Dividable.is-factor d ⟩ m ∎ where open ≡-Reasoning -... | suc t | _ = s≤s z≤n +... | suc t = s≤s z≤n div→k≤m : { m k : ℕ } → k > 1 → m > 0 → Dividable k m → m ≥ k div→k≤m {m} {k} k>1 m>0 d with <-cmp m k diff -r d1b6fb58aad0 -r fa52d72f4bb3 src/ordinal.agda --- a/src/ordinal.agda Sat Aug 26 10:36:09 2023 +0900 +++ b/src/ordinal.agda Mon Jan 01 18:21:36 2024 +0900 @@ -50,11 +50,11 @@ o∅ : {n : Level} → Ordinal {n} o∅ = record { lv = Zero ; ord = Φ Zero } -open import Relation.Binary.HeterogeneousEquality using (_≅_;refl) +-- open import Relation.Binary.HeterogeneousEquality using (_≅_;refl) -ordinal-cong : {n : Level} {x y : Ordinal {n}} → - lv x ≡ lv y → ord x ≅ ord y → x ≡ y -ordinal-cong refl refl = refl +-- ordinal-cong : {n : Level} {x y : Ordinal {n}} → +-- lv x ≡ lv y → ord x ≅ ord y → x ≡ y +-- ordinal-cong refl refl = refl ≡→¬d< : {n : Level} → {lv : ℕ} → {x : OrdinalD {n} lv } → x d< x → ⊥ ≡→¬d< {n} {lx} {OSuc lx y} (s< t) = ≡→¬d< t diff -r d1b6fb58aad0 -r fa52d72f4bb3 src/zf.agda --- a/src/zf.agda Sat Aug 26 10:36:09 2023 +0900 +++ b/src/zf.agda Mon Jan 01 18:21:36 2024 +0900 @@ -1,3 +1,4 @@ +{-# OPTIONS --cubical-compatible --safe #-} module zf where open import Level