Mercurial > hg > Members > kono > Proof > ZF-in-agda
changeset 431:a5f8084b8368
reorganiztion for apkg
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Mon, 21 Dec 2020 10:23:37 +0900 |
parents | 28c7be8f252c |
children | 7476a22edf7e |
files | BAlgbra.agda LEMC.agda LICENSE LICENSE.md OD.agda ODC.agda ODUtil.agda OPair.agda OrdUtil.agda Ordinals.agda README.md Topology.agda VL.agda ZF.agda-lib ZF.agda-pkg cardinal.agda fig/ODandOrdinals.pdf fig/Sets.pdf fig/Sets.txt fig/axiom-dependency.pdf fig/axiom-type.pdf fig/ord-od-mapping.pdf fig/set-theory.pdf fig/zf-record.html fig/zf-record.ind filter.agda generic-filter.agda logic.agda nat.agda ordinal.agda partfunc.agda src/BAlgbra.agda src/LEMC.agda src/OD.agda src/ODC.agda src/ODUtil.agda src/OPair.agda src/OrdUtil.agda src/Ordinals.agda src/Topology.agda src/VL.agda src/cardinal.agda src/filter.agda src/generic-filter.agda src/logic.agda src/nat.agda src/ordinal.agda src/partfunc.agda src/zf.agda src/zfc.agda zf.agda zfc.agda |
diffstat | 52 files changed, 3406 insertions(+), 3937 deletions(-) [+] |
line wrap: on
line diff
--- a/BAlgbra.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,122 +0,0 @@ -open import Level -open import Ordinals -module BAlgbra {n : Level } (O : Ordinals {n}) where - -open import zf -open import logic -import OrdUtil -import OD -import ODUtil -import ODC - -open import Relation.Nullary -open import Relation.Binary -open import Data.Empty -open import Relation.Binary -open import Relation.Binary.Core -open import Relation.Binary.PropositionalEquality -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ; _+_ to _n+_ ) - -open inOrdinal O -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -open OrdUtil O -open ODUtil O - -open OD O -open OD.OD -open ODAxiom odAxiom -open HOD - -open _∧_ -open _∨_ -open Bool - ---_∩_ : ( A B : HOD ) → HOD ---A ∩ B = record { od = record { def = λ x → odef A x ∧ odef B x } ; --- odmax = omin (odmax A) (odmax B) ; <odmax = λ y → min1 (<odmax A (proj1 y)) (<odmax B (proj2 y)) } - -_∪_ : ( A B : HOD ) → HOD -A ∪ B = record { od = record { def = λ x → odef A x ∨ odef B x } ; - odmax = omax (odmax A) (odmax B) ; <odmax = lemma } where - lemma : {y : Ordinal} → odef A y ∨ odef B y → y o< omax (odmax A) (odmax B) - lemma {y} (case1 a) = ordtrans (<odmax A a) (omax-x _ _) - lemma {y} (case2 b) = ordtrans (<odmax B b) (omax-y _ _) - -_\_ : ( A B : HOD ) → HOD -A \ B = record { od = record { def = λ x → odef A x ∧ ( ¬ ( odef B x ) ) }; odmax = odmax A ; <odmax = λ y → <odmax A (proj1 y) } - -∪-Union : { A B : HOD } → Union (A , B) ≡ ( A ∪ B ) -∪-Union {A} {B} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where - lemma1 : {x : Ordinal} → odef (Union (A , B)) x → odef (A ∪ B) x - lemma1 {x} lt = lemma3 lt where - lemma4 : {y : Ordinal} → odef (A , B) y ∧ odef (* y) x → ¬ (¬ ( odef A x ∨ odef B x) ) - lemma4 {y} z with proj1 z - lemma4 {y} z | case1 refl = double-neg (case1 ( subst (λ k → odef k x ) *iso (proj2 z)) ) - lemma4 {y} z | case2 refl = double-neg (case2 ( subst (λ k → odef k x ) *iso (proj2 z)) ) - lemma3 : (((u : Ordinal ) → ¬ odef (A , B) u ∧ odef (* u) x) → ⊥) → odef (A ∪ B) x - lemma3 not = ODC.double-neg-eilm O (FExists _ lemma4 not) -- choice - lemma2 : {x : Ordinal} → odef (A ∪ B) x → odef (Union (A , B)) x - lemma2 {x} (case1 A∋x) = subst (λ k → odef (Union (A , B)) k) &iso ( IsZF.union→ isZF (A , B) (* x) A - ⟪ case1 refl , d→∋ A A∋x ⟫ ) - lemma2 {x} (case2 B∋x) = subst (λ k → odef (Union (A , B)) k) &iso ( IsZF.union→ isZF (A , B) (* x) B - ⟪ case2 refl , d→∋ B B∋x ⟫ ) - -∩-Select : { A B : HOD } → Select A ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) ≡ ( A ∩ B ) -∩-Select {A} {B} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where - lemma1 : {x : Ordinal} → odef (Select A (λ x₁ → (A ∋ x₁) ∧ (B ∋ x₁))) x → odef (A ∩ B) x - lemma1 {x} lt = ⟪ proj1 lt , subst (λ k → odef B k ) &iso (proj2 (proj2 lt)) ⟫ - lemma2 : {x : Ordinal} → odef (A ∩ B) x → odef (Select A (λ x₁ → (A ∋ x₁) ∧ (B ∋ x₁))) x - lemma2 {x} lt = ⟪ proj1 lt , ⟪ d→∋ A (proj1 lt) , d→∋ B (proj2 lt) ⟫ ⟫ - -dist-ord : {p q r : HOD } → p ∩ ( q ∪ r ) ≡ ( p ∩ q ) ∪ ( p ∩ r ) -dist-ord {p} {q} {r} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where - lemma1 : {x : Ordinal} → odef (p ∩ (q ∪ r)) x → odef ((p ∩ q) ∪ (p ∩ r)) x - lemma1 {x} lt with proj2 lt - lemma1 {x} lt | case1 q∋x = case1 ⟪ proj1 lt , q∋x ⟫ - lemma1 {x} lt | case2 r∋x = case2 ⟪ proj1 lt , r∋x ⟫ - lemma2 : {x : Ordinal} → odef ((p ∩ q) ∪ (p ∩ r)) x → odef (p ∩ (q ∪ r)) x - lemma2 {x} (case1 p∩q) = ⟪ proj1 p∩q , case1 (proj2 p∩q ) ⟫ - lemma2 {x} (case2 p∩r) = ⟪ proj1 p∩r , case2 (proj2 p∩r ) ⟫ - -dist-ord2 : {p q r : HOD } → p ∪ ( q ∩ r ) ≡ ( p ∪ q ) ∩ ( p ∪ r ) -dist-ord2 {p} {q} {r} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where - lemma1 : {x : Ordinal} → odef (p ∪ (q ∩ r)) x → odef ((p ∪ q) ∩ (p ∪ r)) x - lemma1 {x} (case1 cp) = ⟪ case1 cp , case1 cp ⟫ - lemma1 {x} (case2 cqr) = ⟪ case2 (proj1 cqr) , case2 (proj2 cqr) ⟫ - lemma2 : {x : Ordinal} → odef ((p ∪ q) ∩ (p ∪ r)) x → odef (p ∪ (q ∩ r)) x - lemma2 {x} lt with proj1 lt | proj2 lt - lemma2 {x} lt | case1 cp | _ = case1 cp - lemma2 {x} lt | _ | case1 cp = case1 cp - lemma2 {x} lt | case2 cq | case2 cr = case2 ⟪ cq , cr ⟫ - -record IsBooleanAlgebra ( L : Set n) - ( b1 : L ) - ( b0 : L ) - ( -_ : L → L ) - ( _+_ : L → L → L ) - ( _x_ : L → L → L ) : Set (suc n) where - field - +-assoc : {a b c : L } → a + ( b + c ) ≡ (a + b) + c - x-assoc : {a b c : L } → a x ( b x c ) ≡ (a x b) x c - +-sym : {a b : L } → a + b ≡ b + a - -sym : {a b : L } → a x b ≡ b x a - +-aab : {a b : L } → a + ( a x b ) ≡ a - x-aab : {a b : L } → a x ( a + b ) ≡ a - +-dist : {a b c : L } → a + ( b x c ) ≡ ( a x b ) + ( a x c ) - x-dist : {a b c : L } → a x ( b + c ) ≡ ( a + b ) x ( a + c ) - a+0 : {a : L } → a + b0 ≡ a - ax1 : {a : L } → a x b1 ≡ a - a+-a1 : {a : L } → a + ( - a ) ≡ b1 - ax-a0 : {a : L } → a x ( - a ) ≡ b0 - -record BooleanAlgebra ( L : Set n) : Set (suc n) where - field - b1 : L - b0 : L - -_ : L → L - _+_ : L → L → L - _x_ : L → L → L - isBooleanAlgebra : IsBooleanAlgebra L b1 b0 -_ _+_ _x_ -
--- a/LEMC.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,178 +0,0 @@ -open import Level -open import Ordinals -open import logic -open import Relation.Nullary -module LEMC {n : Level } (O : Ordinals {n} ) (p∨¬p : ( p : Set n) → p ∨ ( ¬ p )) where - -open import zf -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -open import Relation.Binary.PropositionalEquality -open import Data.Nat.Properties -open import Data.Empty -open import Relation.Binary -open import Relation.Binary.Core - -open import nat -import OD - -open inOrdinal O -open OD O -open OD.OD -open OD._==_ -open ODAxiom odAxiom -import OrdUtil -import ODUtil -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -open OrdUtil O -open ODUtil O - -open import zfc - -open HOD - -open _⊆_ - -decp : ( p : Set n ) → Dec p -- assuming axiom of choice -decp p with p∨¬p p -decp p | case1 x = yes x -decp p | case2 x = no x - -∋-p : (A x : HOD ) → Dec ( A ∋ x ) -∋-p A x with p∨¬p ( A ∋ x) -- LEM -∋-p A x | case1 t = yes t -∋-p A x | case2 t = no (λ x → t x) - -double-neg-eilm : {A : Set n} → ¬ ¬ A → A -- we don't have this in intutionistic logic -double-neg-eilm {A} notnot with decp A -- assuming axiom of choice -... | yes p = p -... | no ¬p = ⊥-elim ( notnot ¬p ) - -power→⊆ : ( A t : HOD) → Power A ∋ t → t ⊆ A -power→⊆ A t PA∋t = record { incl = λ {x} t∋x → double-neg-eilm (λ not → t1 t∋x (λ x → not x) ) } where - t1 : {x : HOD } → t ∋ x → ¬ ¬ (A ∋ x) - t1 = power→ A t PA∋t - ---- With assuption of HOD is ordered, p ∨ ( ¬ p ) <=> axiom of choice ---- -record choiced ( X : Ordinal ) : Set n where - field - a-choice : Ordinal - is-in : odef (* X) a-choice - -open choiced - --- ∋→d : ( a : HOD ) { x : HOD } → * (& a) ∋ x → X ∋ * (a-choice (choice-func X not)) --- ∋→d a lt = subst₂ (λ j k → odef j k) *iso (sym &iso) lt - -oo∋ : { a : HOD} { x : Ordinal } → odef (* (& a)) x → a ∋ * x -oo∋ lt = subst₂ (λ j k → odef j k) *iso (sym &iso) lt - -∋oo : { a : HOD} { x : Ordinal } → a ∋ * x → odef (* (& a)) x -∋oo lt = subst₂ (λ j k → odef j k ) (sym *iso) &iso lt - -OD→ZFC : ZFC -OD→ZFC = record { - ZFSet = HOD - ; _∋_ = _∋_ - ; _≈_ = _=h=_ - ; ∅ = od∅ - ; Select = Select - ; isZFC = isZFC - } where - -- infixr 200 _∈_ - -- infixr 230 _∩_ _∪_ - isZFC : IsZFC (HOD ) _∋_ _=h=_ od∅ Select - isZFC = record { - choice-func = λ A {X} not A∋X → * (a-choice (choice-func X not) ); - choice = λ A {X} A∋X not → oo∋ (is-in (choice-func X not)) - } where - -- - -- the axiom choice from LEM and OD ordering - -- - choice-func : (X : HOD ) → ¬ ( X =h= od∅ ) → choiced (& X) - choice-func X not = have_to_find where - ψ : ( ox : Ordinal ) → Set n - ψ ox = (( x : Ordinal ) → x o< ox → ( ¬ odef X x )) ∨ choiced (& X) - lemma-ord : ( ox : Ordinal ) → ψ ox - lemma-ord ox = TransFinite0 {ψ} induction ox where - ∀-imply-or : {A : Ordinal → Set n } {B : Set n } - → ((x : Ordinal ) → A x ∨ B) → ((x : Ordinal ) → A x) ∨ B - ∀-imply-or {A} {B} ∀AB with p∨¬p ((x : Ordinal ) → A x) -- LEM - ∀-imply-or {A} {B} ∀AB | case1 t = case1 t - ∀-imply-or {A} {B} ∀AB | case2 x = case2 (lemma (λ not → x not )) where - lemma : ¬ ((x : Ordinal ) → A x) → B - lemma not with p∨¬p B - lemma not | case1 b = b - lemma not | case2 ¬b = ⊥-elim (not (λ x → dont-orb (∀AB x) ¬b )) - induction : (x : Ordinal) → ((y : Ordinal) → y o< x → ψ y) → ψ x - induction x prev with ∋-p X ( * x) - ... | yes p = case2 ( record { a-choice = x ; is-in = ∋oo p } ) - ... | no ¬p = lemma where - lemma1 : (y : Ordinal) → (y o< x → odef X y → ⊥) ∨ choiced (& X) - lemma1 y with ∋-p X (* y) - lemma1 y | yes y<X = case2 ( record { a-choice = y ; is-in = ∋oo y<X } ) - lemma1 y | no ¬y<X = case1 ( λ lt y<X → ¬y<X (d→∋ X y<X) ) - lemma : ((y : Ordinal) → y o< x → odef X y → ⊥) ∨ choiced (& X) - lemma = ∀-imply-or lemma1 - odef→o< : {X : HOD } → {x : Ordinal } → odef X x → x o< & X - odef→o< {X} {x} lt = o<-subst {_} {_} {x} {& X} ( c<→o< ( subst₂ (λ j k → odef j k ) (sym *iso) (sym &iso) lt )) &iso &iso - have_to_find : choiced (& X) - have_to_find = dont-or (lemma-ord (& X )) ¬¬X∋x where - ¬¬X∋x : ¬ ((x : Ordinal) → x o< (& X) → odef X x → ⊥) - ¬¬X∋x nn = not record { - eq→ = λ {x} lt → ⊥-elim (nn x (odef→o< lt) lt) - ; eq← = λ {x} lt → ⊥-elim ( ¬x<0 lt ) - } - - -- - -- axiom regurality from ε-induction (using axiom of choice above) - -- - -- from https://math.stackexchange.com/questions/2973777/is-it-possible-to-prove-regularity-with-transfinite-induction-only - -- - -- FIXME : don't use HOD make this level n, so we can remove ε-induction1 - record Minimal (x : HOD) : Set (suc n) where - field - min : HOD - x∋min : x ∋ min - min-empty : (y : HOD ) → ¬ ( min ∋ y) ∧ (x ∋ y) - open Minimal - open _∧_ - induction : {x : HOD} → ({y : HOD} → x ∋ y → (u : HOD ) → (u∋x : u ∋ y) → Minimal u ) - → (u : HOD ) → (u∋x : u ∋ x) → Minimal u - induction {x} prev u u∋x with p∨¬p ((y : Ordinal ) → ¬ (odef x y) ∧ (odef u y)) - ... | case1 P = - record { min = x - ; x∋min = u∋x - ; min-empty = λ y → P (& y) - } - ... | case2 NP = min2 where - p : HOD - p = record { od = record { def = λ y1 → odef x y1 ∧ odef u y1 } ; odmax = omin (odmax x) (odmax u) ; <odmax = lemma } where - lemma : {y : Ordinal} → OD.def (od x) y ∧ OD.def (od u) y → y o< omin (odmax x) (odmax u) - lemma {y} lt = min1 (<odmax x (proj1 lt)) (<odmax u (proj2 lt)) - np : ¬ (p =h= od∅) - np p∅ = NP (λ y p∋y → ∅< {p} {_} (d→∋ p p∋y) p∅ ) - y1choice : choiced (& p) - y1choice = choice-func p np - y1 : HOD - y1 = * (a-choice y1choice) - y1prop : (x ∋ y1) ∧ (u ∋ y1) - y1prop = oo∋ (is-in y1choice) - min2 : Minimal u - min2 = prev (proj1 y1prop) u (proj2 y1prop) - Min2 : (x : HOD) → (u : HOD ) → (u∋x : u ∋ x) → Minimal u - Min2 x u u∋x = (ε-induction {λ y → (u : HOD ) → (u∋x : u ∋ y) → Minimal u } induction x u u∋x ) - cx : {x : HOD} → ¬ (x =h= od∅ ) → choiced (& x ) - cx {x} nx = choice-func x nx - minimal : (x : HOD ) → ¬ (x =h= od∅ ) → HOD - minimal x ne = min (Min2 (* (a-choice (cx {x} ne) )) x ( oo∋ (is-in (cx ne))) ) - x∋minimal : (x : HOD ) → ( ne : ¬ (x =h= od∅ ) ) → odef x ( & ( minimal x ne ) ) - x∋minimal x ne = x∋min (Min2 (* (a-choice (cx {x} ne) )) x ( oo∋ (is-in (cx ne))) ) - minimal-1 : (x : HOD ) → ( ne : ¬ (x =h= od∅ ) ) → (y : HOD ) → ¬ ( odef (minimal x ne) (& y)) ∧ (odef x (& y) ) - minimal-1 x ne y = min-empty (Min2 (* (a-choice (cx ne) )) x ( oo∋ (is-in (cx ne)))) y - - - -
--- a/LICENSE Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,674 +0,0 @@ - GNU GENERAL PUBLIC LICENSE - Version 3, 29 June 2007 - - Copyright (C) 2007 Free Software Foundation, Inc. <https://fsf.org/> - Everyone is permitted to copy and distribute verbatim copies - of this license document, but changing it is not allowed. - - Preamble - - The GNU General Public License is a free, copyleft license for -software and other kinds of works. - - The licenses for most software and other practical works are designed -to take away your freedom to share and change the works. By contrast, -the GNU General Public License is intended to guarantee your freedom to -share and change all versions of a program--to make sure it remains free -software for all its users. We, the Free Software Foundation, use the -GNU General Public License for most of our software; it applies also to -any other work released this way by its authors. You can apply it to -your programs, too. - - When we speak of free software, we are referring to freedom, not -price. Our General Public Licenses are designed to make sure that you -have the freedom to distribute copies of free software (and charge for -them if you wish), that you receive source code or can get it if you -want it, that you can change the software or use pieces of it in new -free programs, and that you know you can do these things. - - To protect your rights, we need to prevent others from denying you -these rights or asking you to surrender the rights. Therefore, you have -certain responsibilities if you distribute copies of the software, or if -you modify it: responsibilities to respect the freedom of others. - - For example, if you distribute copies of such a program, whether -gratis or for a fee, you must pass on to the recipients the same -freedoms that you received. You must make sure that they, too, receive -or can get the source code. And you must show them these terms so they -know their rights. - - Developers that use the GNU GPL protect your rights with two steps: -(1) assert copyright on the software, and (2) offer you this License -giving you legal permission to copy, distribute and/or modify it. - - For the developers' and authors' protection, the GPL clearly explains -that there is no warranty for this free software. For both users' and -authors' sake, the GPL requires that modified versions be marked as -changed, so that their problems will not be attributed erroneously to -authors of previous versions. - - Some devices are designed to deny users access to install or run -modified versions of the software inside them, although the manufacturer -can do so. This is fundamentally incompatible with the aim of -protecting users' freedom to change the software. The systematic -pattern of such abuse occurs in the area of products for individuals to -use, which is precisely where it is most unacceptable. Therefore, we -have designed this version of the GPL to prohibit the practice for those -products. If such problems arise substantially in other domains, we -stand ready to extend this provision to those domains in future versions -of the GPL, as needed to protect the freedom of users. - - Finally, every program is threatened constantly by software patents. -States should not allow patents to restrict development and use of -software on general-purpose computers, but in those that do, we wish to -avoid the special danger that patents applied to a free program could -make it effectively proprietary. To prevent this, the GPL assures that -patents cannot be used to render the program non-free. - - The precise terms and conditions for copying, distribution and -modification follow. - - TERMS AND CONDITIONS - - 0. Definitions. - - "This License" refers to version 3 of the GNU General Public License. - - "Copyright" also means copyright-like laws that apply to other kinds of -works, such as semiconductor masks. - - "The Program" refers to any copyrightable work licensed under this -License. Each licensee is addressed as "you". "Licensees" and -"recipients" may be individuals or organizations. - - To "modify" a work means to copy from or adapt all or part of the work -in a fashion requiring copyright permission, other than the making of an -exact copy. The resulting work is called a "modified version" of the -earlier work or a work "based on" the earlier work. - - A "covered work" means either the unmodified Program or a work based -on the Program. - - To "propagate" a work means to do anything with it that, without -permission, would make you directly or secondarily liable for -infringement under applicable copyright law, except executing it on a -computer or modifying a private copy. Propagation includes copying, -distribution (with or without modification), making available to the -public, and in some countries other activities as well. - - To "convey" a work means any kind of propagation that enables other -parties to make or receive copies. Mere interaction with a user through -a computer network, with no transfer of a copy, is not conveying. - - An interactive user interface displays "Appropriate Legal Notices" -to the extent that it includes a convenient and prominently visible -feature that (1) displays an appropriate copyright notice, and (2) -tells the user that there is no warranty for the work (except to the -extent that warranties are provided), that licensees may convey the -work under this License, and how to view a copy of this License. If -the interface presents a list of user commands or options, such as a -menu, a prominent item in the list meets this criterion. - - 1. Source Code. - - The "source code" for a work means the preferred form of the work -for making modifications to it. "Object code" means any non-source -form of a work. - - A "Standard Interface" means an interface that either is an official -standard defined by a recognized standards body, or, in the case of -interfaces specified for a particular programming language, one that -is widely used among developers working in that language. - - The "System Libraries" of an executable work include anything, other -than the work as a whole, that (a) is included in the normal form of -packaging a Major Component, but which is not part of that Major -Component, and (b) serves only to enable use of the work with that -Major Component, or to implement a Standard Interface for which an -implementation is available to the public in source code form. A -"Major Component", in this context, means a major essential component -(kernel, window system, and so on) of the specific operating system -(if any) on which the executable work runs, or a compiler used to -produce the work, or an object code interpreter used to run it. - - The "Corresponding Source" for a work in object code form means all -the source code needed to generate, install, and (for an executable -work) run the object code and to modify the work, including scripts to -control those activities. However, it does not include the work's -System Libraries, or general-purpose tools or generally available free -programs which are used unmodified in performing those activities but -which are not part of the work. For example, Corresponding Source -includes interface definition files associated with source files for -the work, and the source code for shared libraries and dynamically -linked subprograms that the work is specifically designed to require, -such as by intimate data communication or control flow between those -subprograms and other parts of the work. - - The Corresponding Source need not include anything that users -can regenerate automatically from other parts of the Corresponding -Source. - - The Corresponding Source for a work in source code form is that -same work. - - 2. Basic Permissions. - - All rights granted under this License are granted for the term of -copyright on the Program, and are irrevocable provided the stated -conditions are met. This License explicitly affirms your unlimited -permission to run the unmodified Program. The output from running a -covered work is covered by this License only if the output, given its -content, constitutes a covered work. This License acknowledges your -rights of fair use or other equivalent, as provided by copyright law. - - You may make, run and propagate covered works that you do not -convey, without conditions so long as your license otherwise remains -in force. You may convey covered works to others for the sole purpose -of having them make modifications exclusively for you, or provide you -with facilities for running those works, provided that you comply with -the terms of this License in conveying all material for which you do -not control copyright. Those thus making or running the covered works -for you must do so exclusively on your behalf, under your direction -and control, on terms that prohibit them from making any copies of -your copyrighted material outside their relationship with you. - - Conveying under any other circumstances is permitted solely under -the conditions stated below. Sublicensing is not allowed; section 10 -makes it unnecessary. - - 3. Protecting Users' Legal Rights From Anti-Circumvention Law. - - No covered work shall be deemed part of an effective technological -measure under any applicable law fulfilling obligations under article -11 of the WIPO copyright treaty adopted on 20 December 1996, or -similar laws prohibiting or restricting circumvention of such -measures. - - When you convey a covered work, you waive any legal power to forbid -circumvention of technological measures to the extent such circumvention -is effected by exercising rights under this License with respect to -the covered work, and you disclaim any intention to limit operation or -modification of the work as a means of enforcing, against the work's -users, your or third parties' legal rights to forbid circumvention of -technological measures. - - 4. Conveying Verbatim Copies. - - You may convey verbatim copies of the Program's source code as you -receive it, in any medium, provided that you conspicuously and -appropriately publish on each copy an appropriate copyright notice; -keep intact all notices stating that this License and any -non-permissive terms added in accord with section 7 apply to the code; -keep intact all notices of the absence of any warranty; and give all -recipients a copy of this License along with the Program. - - You may charge any price or no price for each copy that you convey, -and you may offer support or warranty protection for a fee. - - 5. Conveying Modified Source Versions. - - You may convey a work based on the Program, or the modifications to -produce it from the Program, in the form of source code under the -terms of section 4, provided that you also meet all of these conditions: - - a) The work must carry prominent notices stating that you modified - it, and giving a relevant date. - - b) The work must carry prominent notices stating that it is - released under this License and any conditions added under section - 7. This requirement modifies the requirement in section 4 to - "keep intact all notices". - - c) You must license the entire work, as a whole, under this - License to anyone who comes into possession of a copy. This - License will therefore apply, along with any applicable section 7 - additional terms, to the whole of the work, and all its parts, - regardless of how they are packaged. This License gives no - permission to license the work in any other way, but it does not - invalidate such permission if you have separately received it. - - d) If the work has interactive user interfaces, each must display - Appropriate Legal Notices; however, if the Program has interactive - interfaces that do not display Appropriate Legal Notices, your - work need not make them do so. - - A compilation of a covered work with other separate and independent -works, which are not by their nature extensions of the covered work, -and which are not combined with it such as to form a larger program, -in or on a volume of a storage or distribution medium, is called an -"aggregate" if the compilation and its resulting copyright are not -used to limit the access or legal rights of the compilation's users -beyond what the individual works permit. Inclusion of a covered work -in an aggregate does not cause this License to apply to the other -parts of the aggregate. - - 6. Conveying Non-Source Forms. - - You may convey a covered work in object code form under the terms -of sections 4 and 5, provided that you also convey the -machine-readable Corresponding Source under the terms of this License, -in one of these ways: - - a) Convey the object code in, or embodied in, a physical product - (including a physical distribution medium), accompanied by the - Corresponding Source fixed on a durable physical medium - customarily used for software interchange. - - b) Convey the object code in, or embodied in, a physical product - (including a physical distribution medium), accompanied by a - written offer, valid for at least three years and valid for as - long as you offer spare parts or customer support for that product - model, to give anyone who possesses the object code either (1) a - copy of the Corresponding Source for all the software in the - product that is covered by this License, on a durable physical - medium customarily used for software interchange, for a price no - more than your reasonable cost of physically performing this - conveying of source, or (2) access to copy the - Corresponding Source from a network server at no charge. - - c) Convey individual copies of the object code with a copy of the - written offer to provide the Corresponding Source. This - alternative is allowed only occasionally and noncommercially, and - only if you received the object code with such an offer, in accord - with subsection 6b. - - d) Convey the object code by offering access from a designated - place (gratis or for a charge), and offer equivalent access to the - Corresponding Source in the same way through the same place at no - further charge. You need not require recipients to copy the - Corresponding Source along with the object code. If the place to - copy the object code is a network server, the Corresponding Source - may be on a different server (operated by you or a third party) - that supports equivalent copying facilities, provided you maintain - clear directions next to the object code saying where to find the - Corresponding Source. Regardless of what server hosts the - Corresponding Source, you remain obligated to ensure that it is - available for as long as needed to satisfy these requirements. - - e) Convey the object code using peer-to-peer transmission, provided - you inform other peers where the object code and Corresponding - Source of the work are being offered to the general public at no - charge under subsection 6d. - - A separable portion of the object code, whose source code is excluded -from the Corresponding Source as a System Library, need not be -included in conveying the object code work. - - A "User Product" is either (1) a "consumer product", which means any -tangible personal property which is normally used for personal, family, -or household purposes, or (2) anything designed or sold for incorporation -into a dwelling. In determining whether a product is a consumer product, -doubtful cases shall be resolved in favor of coverage. For a particular -product received by a particular user, "normally used" refers to a -typical or common use of that class of product, regardless of the status -of the particular user or of the way in which the particular user -actually uses, or expects or is expected to use, the product. A product -is a consumer product regardless of whether the product has substantial -commercial, industrial or non-consumer uses, unless such uses represent -the only significant mode of use of the product. - - "Installation Information" for a User Product means any methods, -procedures, authorization keys, or other information required to install -and execute modified versions of a covered work in that User Product from -a modified version of its Corresponding Source. The information must -suffice to ensure that the continued functioning of the modified object -code is in no case prevented or interfered with solely because -modification has been made. - - If you convey an object code work under this section in, or with, or -specifically for use in, a User Product, and the conveying occurs as -part of a transaction in which the right of possession and use of the -User Product is transferred to the recipient in perpetuity or for a -fixed term (regardless of how the transaction is characterized), the -Corresponding Source conveyed under this section must be accompanied -by the Installation Information. But this requirement does not apply -if neither you nor any third party retains the ability to install -modified object code on the User Product (for example, the work has -been installed in ROM). - - The requirement to provide Installation Information does not include a -requirement to continue to provide support service, warranty, or updates -for a work that has been modified or installed by the recipient, or for -the User Product in which it has been modified or installed. Access to a -network may be denied when the modification itself materially and -adversely affects the operation of the network or violates the rules and -protocols for communication across the network. - - Corresponding Source conveyed, and Installation Information provided, -in accord with this section must be in a format that is publicly -documented (and with an implementation available to the public in -source code form), and must require no special password or key for -unpacking, reading or copying. - - 7. Additional Terms. - - "Additional permissions" are terms that supplement the terms of this -License by making exceptions from one or more of its conditions. -Additional permissions that are applicable to the entire Program shall -be treated as though they were included in this License, to the extent -that they are valid under applicable law. If additional permissions -apply only to part of the Program, that part may be used separately -under those permissions, but the entire Program remains governed by -this License without regard to the additional permissions. - - When you convey a copy of a covered work, you may at your option -remove any additional permissions from that copy, or from any part of -it. (Additional permissions may be written to require their own -removal in certain cases when you modify the work.) You may place -additional permissions on material, added by you to a covered work, -for which you have or can give appropriate copyright permission. - - Notwithstanding any other provision of this License, for material you -add to a covered work, you may (if authorized by the copyright holders of -that material) supplement the terms of this License with terms: - - a) Disclaiming warranty or limiting liability differently from the - terms of sections 15 and 16 of this License; or - - b) Requiring preservation of specified reasonable legal notices or - author attributions in that material or in the Appropriate Legal - Notices displayed by works containing it; or - - c) Prohibiting misrepresentation of the origin of that material, or - requiring that modified versions of such material be marked in - reasonable ways as different from the original version; or - - d) Limiting the use for publicity purposes of names of licensors or - authors of the material; or - - e) Declining to grant rights under trademark law for use of some - trade names, trademarks, or service marks; or - - f) Requiring indemnification of licensors and authors of that - material by anyone who conveys the material (or modified versions of - it) with contractual assumptions of liability to the recipient, for - any liability that these contractual assumptions directly impose on - those licensors and authors. - - All other non-permissive additional terms are considered "further -restrictions" within the meaning of section 10. If the Program as you -received it, or any part of it, contains a notice stating that it is -governed by this License along with a term that is a further -restriction, you may remove that term. If a license document contains -a further restriction but permits relicensing or conveying under this -License, you may add to a covered work material governed by the terms -of that license document, provided that the further restriction does -not survive such relicensing or conveying. - - If you add terms to a covered work in accord with this section, you -must place, in the relevant source files, a statement of the -additional terms that apply to those files, or a notice indicating -where to find the applicable terms. - - Additional terms, permissive or non-permissive, may be stated in the -form of a separately written license, or stated as exceptions; -the above requirements apply either way. - - 8. Termination. - - You may not propagate or modify a covered work except as expressly -provided under this License. Any attempt otherwise to propagate or -modify it is void, and will automatically terminate your rights under -this License (including any patent licenses granted under the third -paragraph of section 11). - - However, if you cease all violation of this License, then your -license from a particular copyright holder is reinstated (a) -provisionally, unless and until the copyright holder explicitly and -finally terminates your license, and (b) permanently, if the copyright -holder fails to notify you of the violation by some reasonable means -prior to 60 days after the cessation. - - Moreover, your license from a particular copyright holder is -reinstated permanently if the copyright holder notifies you of the -violation by some reasonable means, this is the first time you have -received notice of violation of this License (for any work) from that -copyright holder, and you cure the violation prior to 30 days after -your receipt of the notice. - - Termination of your rights under this section does not terminate the -licenses of parties who have received copies or rights from you under -this License. If your rights have been terminated and not permanently -reinstated, you do not qualify to receive new licenses for the same -material under section 10. - - 9. Acceptance Not Required for Having Copies. - - You are not required to accept this License in order to receive or -run a copy of the Program. Ancillary propagation of a covered work -occurring solely as a consequence of using peer-to-peer transmission -to receive a copy likewise does not require acceptance. However, -nothing other than this License grants you permission to propagate or -modify any covered work. These actions infringe copyright if you do -not accept this License. Therefore, by modifying or propagating a -covered work, you indicate your acceptance of this License to do so. - - 10. Automatic Licensing of Downstream Recipients. - - Each time you convey a covered work, the recipient automatically -receives a license from the original licensors, to run, modify and -propagate that work, subject to this License. You are not responsible -for enforcing compliance by third parties with this License. - - An "entity transaction" is a transaction transferring control of an -organization, or substantially all assets of one, or subdividing an -organization, or merging organizations. If propagation of a covered -work results from an entity transaction, each party to that -transaction who receives a copy of the work also receives whatever -licenses to the work the party's predecessor in interest had or could -give under the previous paragraph, plus a right to possession of the -Corresponding Source of the work from the predecessor in interest, if -the predecessor has it or can get it with reasonable efforts. - - You may not impose any further restrictions on the exercise of the -rights granted or affirmed under this License. For example, you may -not impose a license fee, royalty, or other charge for exercise of -rights granted under this License, and you may not initiate litigation -(including a cross-claim or counterclaim in a lawsuit) alleging that -any patent claim is infringed by making, using, selling, offering for -sale, or importing the Program or any portion of it. - - 11. Patents. - - A "contributor" is a copyright holder who authorizes use under this -License of the Program or a work on which the Program is based. The -work thus licensed is called the contributor's "contributor version". - - A contributor's "essential patent claims" are all patent claims -owned or controlled by the contributor, whether already acquired or -hereafter acquired, that would be infringed by some manner, permitted -by this License, of making, using, or selling its contributor version, -but do not include claims that would be infringed only as a -consequence of further modification of the contributor version. For -purposes of this definition, "control" includes the right to grant -patent sublicenses in a manner consistent with the requirements of -this License. - - Each contributor grants you a non-exclusive, worldwide, royalty-free -patent license under the contributor's essential patent claims, to -make, use, sell, offer for sale, import and otherwise run, modify and -propagate the contents of its contributor version. - - In the following three paragraphs, a "patent license" is any express -agreement or commitment, however denominated, not to enforce a patent -(such as an express permission to practice a patent or covenant not to -sue for patent infringement). To "grant" such a patent license to a -party means to make such an agreement or commitment not to enforce a -patent against the party. - - If you convey a covered work, knowingly relying on a patent license, -and the Corresponding Source of the work is not available for anyone -to copy, free of charge and under the terms of this License, through a -publicly available network server or other readily accessible means, -then you must either (1) cause the Corresponding Source to be so -available, or (2) arrange to deprive yourself of the benefit of the -patent license for this particular work, or (3) arrange, in a manner -consistent with the requirements of this License, to extend the patent -license to downstream recipients. "Knowingly relying" means you have -actual knowledge that, but for the patent license, your conveying the -covered work in a country, or your recipient's use of the covered work -in a country, would infringe one or more identifiable patents in that -country that you have reason to believe are valid. - - If, pursuant to or in connection with a single transaction or -arrangement, you convey, or propagate by procuring conveyance of, a -covered work, and grant a patent license to some of the parties -receiving the covered work authorizing them to use, propagate, modify -or convey a specific copy of the covered work, then the patent license -you grant is automatically extended to all recipients of the covered -work and works based on it. - - A patent license is "discriminatory" if it does not include within -the scope of its coverage, prohibits the exercise of, or is -conditioned on the non-exercise of one or more of the rights that are -specifically granted under this License. You may not convey a covered -work if you are a party to an arrangement with a third party that is -in the business of distributing software, under which you make payment -to the third party based on the extent of your activity of conveying -the work, and under which the third party grants, to any of the -parties who would receive the covered work from you, a discriminatory -patent license (a) in connection with copies of the covered work -conveyed by you (or copies made from those copies), or (b) primarily -for and in connection with specific products or compilations that -contain the covered work, unless you entered into that arrangement, -or that patent license was granted, prior to 28 March 2007. - - Nothing in this License shall be construed as excluding or limiting -any implied license or other defenses to infringement that may -otherwise be available to you under applicable patent law. - - 12. No Surrender of Others' Freedom. - - If conditions are imposed on you (whether by court order, agreement or -otherwise) that contradict the conditions of this License, they do not -excuse you from the conditions of this License. If you cannot convey a -covered work so as to satisfy simultaneously your obligations under this -License and any other pertinent obligations, then as a consequence you may -not convey it at all. For example, if you agree to terms that obligate you -to collect a royalty for further conveying from those to whom you convey -the Program, the only way you could satisfy both those terms and this -License would be to refrain entirely from conveying the Program. - - 13. Use with the GNU Affero General Public License. - - Notwithstanding any other provision of this License, you have -permission to link or combine any covered work with a work licensed -under version 3 of the GNU Affero General Public License into a single -combined work, and to convey the resulting work. The terms of this -License will continue to apply to the part which is the covered work, -but the special requirements of the GNU Affero General Public License, -section 13, concerning interaction through a network will apply to the -combination as such. - - 14. Revised Versions of this License. - - The Free Software Foundation may publish revised and/or new versions of -the GNU General Public License from time to time. Such new versions will -be similar in spirit to the present version, but may differ in detail to -address new problems or concerns. - - Each version is given a distinguishing version number. If the -Program specifies that a certain numbered version of the GNU General -Public License "or any later version" applies to it, you have the -option of following the terms and conditions either of that numbered -version or of any later version published by the Free Software -Foundation. If the Program does not specify a version number of the -GNU General Public License, you may choose any version ever published -by the Free Software Foundation. - - If the Program specifies that a proxy can decide which future -versions of the GNU General Public License can be used, that proxy's -public statement of acceptance of a version permanently authorizes you -to choose that version for the Program. - - Later license versions may give you additional or different -permissions. However, no additional obligations are imposed on any -author or copyright holder as a result of your choosing to follow a -later version. - - 15. Disclaimer of Warranty. - - THERE IS NO WARRANTY FOR THE PROGRAM, TO THE EXTENT PERMITTED BY -APPLICABLE LAW. EXCEPT WHEN OTHERWISE STATED IN WRITING THE COPYRIGHT -HOLDERS AND/OR OTHER PARTIES PROVIDE THE PROGRAM "AS IS" WITHOUT WARRANTY -OF ANY KIND, EITHER EXPRESSED OR IMPLIED, INCLUDING, BUT NOT LIMITED TO, -THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR -PURPOSE. THE ENTIRE RISK AS TO THE QUALITY AND PERFORMANCE OF THE PROGRAM -IS WITH YOU. SHOULD THE PROGRAM PROVE DEFECTIVE, YOU ASSUME THE COST OF -ALL NECESSARY SERVICING, REPAIR OR CORRECTION. - - 16. Limitation of Liability. - - IN NO EVENT UNLESS REQUIRED BY APPLICABLE LAW OR AGREED TO IN WRITING -WILL ANY COPYRIGHT HOLDER, OR ANY OTHER PARTY WHO MODIFIES AND/OR CONVEYS -THE PROGRAM AS PERMITTED ABOVE, BE LIABLE TO YOU FOR DAMAGES, INCLUDING ANY -GENERAL, SPECIAL, INCIDENTAL OR CONSEQUENTIAL DAMAGES ARISING OUT OF THE -USE OR INABILITY TO USE THE PROGRAM (INCLUDING BUT NOT LIMITED TO LOSS OF -DATA OR DATA BEING RENDERED INACCURATE OR LOSSES SUSTAINED BY YOU OR THIRD -PARTIES OR A FAILURE OF THE PROGRAM TO OPERATE WITH ANY OTHER PROGRAMS), -EVEN IF SUCH HOLDER OR OTHER PARTY HAS BEEN ADVISED OF THE POSSIBILITY OF -SUCH DAMAGES. - - 17. Interpretation of Sections 15 and 16. - - If the disclaimer of warranty and limitation of liability provided -above cannot be given local legal effect according to their terms, -reviewing courts shall apply local law that most closely approximates -an absolute waiver of all civil liability in connection with the -Program, unless a warranty or assumption of liability accompanies a -copy of the Program in return for a fee. - - END OF TERMS AND CONDITIONS - - How to Apply These Terms to Your New Programs - - If you develop a new program, and you want it to be of the greatest -possible use to the public, the best way to achieve this is to make it -free software which everyone can redistribute and change under these terms. - - To do so, attach the following notices to the program. It is safest -to attach them to the start of each source file to most effectively -state the exclusion of warranty; and each file should have at least -the "copyright" line and a pointer to where the full notice is found. - - <one line to give the program's name and a brief idea of what it does.> - Copyright (C) <year> <name of author> - - This program is free software: you can redistribute it and/or modify - it under the terms of the GNU General Public License as published by - the Free Software Foundation, either version 3 of the License, or - (at your option) any later version. - - This program is distributed in the hope that it will be useful, - but WITHOUT ANY WARRANTY; without even the implied warranty of - MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the - GNU General Public License for more details. - - You should have received a copy of the GNU General Public License - along with this program. If not, see <https://www.gnu.org/licenses/>. - -Also add information on how to contact you by electronic and paper mail. - - If the program does terminal interaction, make it output a short -notice like this when it starts in an interactive mode: - - <program> Copyright (C) <year> <name of author> - This program comes with ABSOLUTELY NO WARRANTY; for details type `show w'. - This is free software, and you are welcome to redistribute it - under certain conditions; type `show c' for details. - -The hypothetical commands `show w' and `show c' should show the appropriate -parts of the General Public License. Of course, your program's commands -might be different; for a GUI interface, you would use an "about box". - - You should also get your employer (if you work as a programmer) or school, -if any, to sign a "copyright disclaimer" for the program, if necessary. -For more information on this, and how to apply and follow the GNU GPL, see -<https://www.gnu.org/licenses/>. - - The GNU General Public License does not permit incorporating your program -into proprietary programs. If your program is a subroutine library, you -may consider it more useful to permit linking proprietary applications with -the library. If this is what you want to do, use the GNU Lesser General -Public License instead of this License. But first, please read -<https://www.gnu.org/licenses/why-not-lgpl.html>.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/LICENSE.md Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,19 @@ +Copyright (c) 2020 <copyright Shinji KONO, University of the Ryukyus, kono@ie.u-ryukyu.ac.jp > + +Permission is hereby granted, free of charge, to any person obtaining a copy +of this software and associated documentation files (the "Software"), to deal +in the Software without restriction, including without limitation the rights +to use, copy, modify, merge, publish, distribute, sublicense, and/or sell +copies of the Software, and to permit persons to whom the Software is +furnished to do so, subject to the following conditions: + +The above copyright notice and this permission notice shall be included in all +copies or substantial portions of the Software. + +THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR +IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, +FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE +AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER +LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, +OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE +SOFTWARE.
--- a/OD.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,618 +0,0 @@ -{-# OPTIONS --allow-unsolved-metas #-} -open import Level -open import Ordinals -module OD {n : Level } (O : Ordinals {n} ) where - -open import zf -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -open import Relation.Binary.PropositionalEquality hiding ( [_] ) -open import Data.Nat.Properties -open import Data.Empty -open import Relation.Nullary -open import Relation.Binary hiding (_⇔_) -open import Relation.Binary.Core hiding (_⇔_) - -open import logic -import OrdUtil -open import nat - -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -open OrdUtil O - --- Ordinal Definable Set - -record OD : Set (suc n ) where - field - def : (x : Ordinal ) → Set n - -open 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 _==_ - -==-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 - --- next assumptions are our axiom --- --- 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. --- --- ==→o≡ is necessary to prove axiom of extensionality. - --- Ordinals in OD , the maximum -Ords : OD -Ords = record { def = λ x → One } - -record HOD : Set (suc n) where - field - od : OD - odmax : Ordinal - <odmax : {y : Ordinal} → def od y → y o< odmax - -open HOD - -record ODAxiom : Set (suc n) where - field - -- HOD is isomorphic to Ordinal (by means of Goedel number) - & : HOD → Ordinal - * : Ordinal → HOD - c<→o< : {x y : HOD } → def (od y) ( & x ) → & x o< & y - ⊆→o≤ : {y z : HOD } → ({x : Ordinal} → def (od y) x → def (od z) x ) → & y o< osuc (& z) - *iso : {x : HOD } → * ( & x ) ≡ x - &iso : {x : Ordinal } → & ( * x ) ≡ x - ==→o≡ : {x y : HOD } → (od x == od y) → x ≡ y - sup-o : (A : HOD) → ( ( x : Ordinal ) → def (od A) x → Ordinal ) → Ordinal - sup-o< : (A : HOD) → { ψ : ( x : Ordinal ) → def (od A) x → Ordinal } → ∀ {x : Ordinal } → (lt : def (od A) x ) → ψ x lt o< sup-o A ψ --- possible order restriction - ho< : {x : HOD} → & x o< next (odmax x) - - -postulate odAxiom : ODAxiom -open ODAxiom odAxiom - --- odmax minimality --- --- since we have ==→o≡ , so odmax have to be unique. We should have odmaxmin in HOD. --- We can calculate the minimum using sup but it is tedius. --- Only Select has non minimum odmax. --- We have the same problem on 'def' itself, but we leave it. - -odmaxmin : Set (suc n) -odmaxmin = (y : HOD) (z : Ordinal) → ((x : Ordinal)→ def (od y) x → x o< z) → odmax y o< osuc z - --- OD ⇔ Ordinal leads a contradiction, so we need bounded HOD -¬OD-order : ( & : OD → Ordinal ) → ( * : Ordinal → OD ) → ( { x y : OD } → def y ( & x ) → & x o< & y) → ⊥ -¬OD-order & * c<→o< = osuc-< <-osuc (c<→o< {Ords} OneObj ) - --- Open supreme upper bound leads a contradition, so we use domain restriction on sup -¬open-sup : ( sup-o : (Ordinal → Ordinal ) → Ordinal) → ((ψ : Ordinal → Ordinal ) → (x : Ordinal) → ψ x o< sup-o ψ ) → ⊥ -¬open-sup sup-o sup-o< = o<> <-osuc (sup-o< next-ord (sup-o next-ord)) where - next-ord : Ordinal → Ordinal - next-ord x = osuc x - --- Ordinal in OD ( and ZFSet ) Transitive Set -Ord : ( a : Ordinal ) → HOD -Ord a = record { od = record { def = λ y → y o< a } ; odmax = a ; <odmax = lemma } where - lemma : {x : Ordinal} → x o< a → x o< a - lemma {x} lt = lt - -od∅ : HOD -od∅ = Ord o∅ - -odef : HOD → Ordinal → Set n -odef A x = def ( od A ) x - -_∋_ : ( a x : HOD ) → Set n -_∋_ a x = odef a ( & x ) - --- _c<_ : ( x a : HOD ) → Set n --- x c< a = a ∋ x - -d→∋ : ( a : HOD ) { x : Ordinal} → odef a x → a ∋ (* x) -d→∋ a lt = subst (λ k → odef a k ) (sym &iso) lt - --- odef-subst : {Z : HOD } {X : Ordinal }{z : HOD } {x : Ordinal }→ odef Z X → Z ≡ z → X ≡ x → odef z x --- odef-subst df refl refl = df - -otrans : {a x y : Ordinal } → odef (Ord a) x → odef (Ord x) y → odef (Ord a) y -otrans x<a y<x = ordtrans y<x x<a - --- If we have reverse of c<→o<, everything becomes Ordinal -∈→c<→HOD=Ord : ( o<→c< : {x y : Ordinal } → x o< y → odef (* y) x ) → {x : HOD } → x ≡ Ord (& x) -∈→c<→HOD=Ord o<→c< {x} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where - lemma1 : {y : Ordinal} → odef x y → odef (Ord (& x)) y - lemma1 {y} lt = subst ( λ k → k o< & x ) &iso (c<→o< {* y} {x} (d→∋ x lt)) - lemma2 : {y : Ordinal} → odef (Ord (& x)) y → odef x y - lemma2 {y} lt = subst (λ k → odef k y ) *iso (o<→c< {y} {& x} lt ) - --- avoiding lv != Zero error -orefl : { x : HOD } → { y : Ordinal } → & x ≡ y → & x ≡ y -orefl refl = refl - -==-iso : { x y : HOD } → od (* (& x)) == od (* (& y)) → od x == od y -==-iso {x} {y} eq = record { - eq→ = λ {z} d → lemma ( eq→ eq (subst (λ k → odef k z ) (sym *iso) d )) ; - eq← = λ {z} d → lemma ( eq← eq (subst (λ k → odef k z ) (sym *iso) d )) } - where - lemma : {x : HOD } {z : Ordinal } → odef (* (& x)) z → odef x z - lemma {x} {z} d = subst (λ k → odef k z) (*iso) d - -=-iso : {x y : HOD } → (od x == od y) ≡ (od (* (& x)) == od y) -=-iso {_} {y} = cong ( λ k → od k == od y ) (sym *iso) - -ord→== : { x y : HOD } → & x ≡ & y → od x == od y -ord→== {x} {y} eq = ==-iso (lemma (& x) (& y) (orefl eq)) where - lemma : ( ox oy : Ordinal ) → ox ≡ oy → od (* ox) == od (* oy) - lemma ox ox refl = ==-refl - -o≡→== : { x y : Ordinal } → x ≡ y → od (* x) == od (* y) -o≡→== {x} {.x} refl = ==-refl - -o∅≡od∅ : * (o∅ ) ≡ od∅ -o∅≡od∅ = ==→o≡ lemma where - lemma0 : {x : Ordinal} → odef (* o∅) x → odef od∅ x - lemma0 {x} lt with c<→o< {* x} {* o∅} (subst (λ k → odef (* o∅) k ) (sym &iso) lt) - ... | t = subst₂ (λ j k → j o< k ) &iso &iso t - lemma1 : {x : Ordinal} → odef od∅ x → odef (* o∅) x - lemma1 {x} lt = ⊥-elim (¬x<0 lt) - lemma : od (* o∅) == od od∅ - lemma = record { eq→ = lemma0 ; eq← = lemma1 } - -ord-od∅ : & (od∅ ) ≡ o∅ -ord-od∅ = sym ( subst (λ k → k ≡ & (od∅ ) ) &iso (cong ( λ k → & k ) o∅≡od∅ ) ) - -≡o∅→=od∅ : {x : HOD} → & x ≡ o∅ → od x == od od∅ -≡o∅→=od∅ {x} eq = record { eq→ = λ {y} lt → ⊥-elim ( ¬x<0 {y} (subst₂ (λ j k → j o< k ) &iso eq ( c<→o< {* y} {x} (d→∋ x lt)))) - ; eq← = λ {y} lt → ⊥-elim ( ¬x<0 lt )} - -=od∅→≡o∅ : {x : HOD} → od x == od od∅ → & x ≡ o∅ -=od∅→≡o∅ {x} eq = trans (cong (λ k → & k ) (==→o≡ {x} {od∅} eq)) ord-od∅ - -∅0 : record { def = λ x → Lift n ⊥ } == od od∅ -eq→ ∅0 {w} (lift ()) -eq← ∅0 {w} lt = lift (¬x<0 lt) - -∅< : { x y : HOD } → odef x (& y ) → ¬ ( od x == od od∅ ) -∅< {x} {y} d eq with eq→ (==-trans eq (==-sym ∅0) ) d -∅< {x} {y} d eq | lift () - -∅6 : { x : HOD } → ¬ ( x ∋ x ) -- no Russel paradox -∅6 {x} x∋x = o<¬≡ refl ( c<→o< {x} {x} x∋x ) - -odef-iso : {A B : HOD } {x y : Ordinal } → x ≡ y → (odef A y → odef B y) → odef A x → odef B x -odef-iso refl t = t - -is-o∅ : ( x : Ordinal ) → Dec ( x ≡ o∅ ) -is-o∅ x with trio< x o∅ -is-o∅ x | tri< a ¬b ¬c = no ¬b -is-o∅ x | tri≈ ¬a b ¬c = yes b -is-o∅ x | tri> ¬a ¬b c = no ¬b - --- the pair -_,_ : HOD → HOD → HOD -x , y = record { od = record { def = λ t → (t ≡ & x ) ∨ ( t ≡ & y ) } ; odmax = omax (& x) (& y) ; <odmax = lemma } where - lemma : {t : Ordinal} → (t ≡ & x) ∨ (t ≡ & y) → t o< omax (& x) (& y) - lemma {t} (case1 refl) = omax-x _ _ - lemma {t} (case2 refl) = omax-y _ _ - -pair<y : {x y : HOD } → y ∋ x → & (x , x) o< osuc (& y) -pair<y {x} {y} y∋x = ⊆→o≤ lemma where - lemma : {z : Ordinal} → def (od (x , x)) z → def (od y) z - lemma (case1 refl) = y∋x - lemma (case2 refl) = y∋x - --- another possible restriction. We reqest no minimality on odmax, so it may arbitrary larger. -odmax<& : { x y : HOD } → x ∋ y → Set n -odmax<& {x} {y} x∋y = odmax x o< & x - -in-codomain : (X : HOD ) → ( ψ : HOD → HOD ) → OD -in-codomain X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( odef X y ∧ ( x ≡ & (ψ (* y ))))) } - -_∩_ : ( A B : HOD ) → HOD -A ∩ B = record { od = record { def = λ x → odef A x ∧ odef B x } - ; odmax = omin (odmax A) (odmax B) ; <odmax = λ y → min1 (<odmax A (proj1 y)) (<odmax B (proj2 y))} - -record _⊆_ ( A B : HOD ) : Set (suc n) where - field - incl : { x : HOD } → A ∋ x → B ∋ x - -open _⊆_ -infixr 220 _⊆_ - --- if we have & (x , x) ≡ osuc (& x), ⊆→o≤ → c<→o< -⊆→o≤→c<→o< : ({x : HOD} → & (x , x) ≡ osuc (& x) ) - → ({y z : HOD } → ({x : Ordinal} → def (od y) x → def (od z) x ) → & y o< osuc (& z) ) - → {x y : HOD } → def (od y) ( & x ) → & x o< & y -⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x with trio< (& x) (& y) -⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri< a ¬b ¬c = a -⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri≈ ¬a b ¬c = ⊥-elim ( o<¬≡ (peq {x}) (pair<y (subst (λ k → k ∋ x) (sym ( ==→o≡ {x} {y} (ord→== b))) y∋x ))) -⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri> ¬a ¬b c = - ⊥-elim ( o<> (⊆→o≤ {x , x} {y} y⊆x,x ) lemma1 ) where - lemma : {z : Ordinal} → (z ≡ & x) ∨ (z ≡ & x) → & x ≡ z - lemma (case1 refl) = refl - lemma (case2 refl) = refl - y⊆x,x : {z : Ordinal} → def (od (x , x)) z → def (od y) z - y⊆x,x {z} lt = subst (λ k → def (od y) k ) (lemma lt) y∋x - lemma1 : osuc (& y) o< & (x , x) - lemma1 = subst (λ k → osuc (& y) o< k ) (sym (peq {x})) (osucc c ) - -ε-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 - -Select : (X : HOD ) → ((x : HOD ) → Set n ) → HOD -Select X ψ = record { od = record { def = λ x → ( odef X x ∧ ψ ( * x )) } ; odmax = odmax X ; <odmax = λ y → <odmax X (proj1 y) } - -Replace : HOD → (HOD → HOD) → HOD -Replace X ψ = record { od = record { def = λ x → (x o< sup-o X (λ y X∋y → & (ψ (* y)))) ∧ def (in-codomain X ψ) x } - ; odmax = rmax ; <odmax = rmax<} where - rmax : Ordinal - rmax = sup-o X (λ y X∋y → & (ψ (* y))) - rmax< : {y : Ordinal} → (y o< rmax) ∧ def (in-codomain X ψ) y → y o< rmax - rmax< lt = proj1 lt - --- --- If we have LEM, Replace' is equivalent to Replace --- -in-codomain' : (X : HOD ) → ((x : HOD) → X ∋ x → HOD) → OD -in-codomain' X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( odef X y ∧ ((lt : odef X y) → x ≡ & (ψ (* y ) (d→∋ X lt) )))) } -Replace' : (X : HOD) → ((x : HOD) → X ∋ x → HOD) → HOD -Replace' X ψ = record { od = record { def = λ x → (x o< sup-o X (λ y X∋y → & (ψ (* y) (d→∋ X X∋y) ))) ∧ def (in-codomain' X ψ) x } - ; odmax = rmax ; <odmax = rmax< } where - rmax : Ordinal - rmax = sup-o X (λ y X∋y → & (ψ (* y) (d→∋ X X∋y))) - rmax< : {y : Ordinal} → (y o< rmax) ∧ def (in-codomain' X ψ) y → y o< rmax - rmax< lt = proj1 lt - -Union : HOD → HOD -Union U = record { od = record { def = λ x → ¬ (∀ (u : Ordinal ) → ¬ ((odef U u) ∧ (odef (* u) x))) } - ; odmax = osuc (& U) ; <odmax = umax< } where - umax< : {y : Ordinal} → ¬ ((u : Ordinal) → ¬ def (od U) u ∧ def (od (* u)) y) → y o< osuc (& U) - umax< {y} not = lemma (FExists _ lemma1 not ) where - lemma0 : {x : Ordinal} → def (od (* x)) y → y o< x - lemma0 {x} x<y = subst₂ (λ j k → j o< k ) &iso &iso (c<→o< (d→∋ (* x) x<y )) - lemma2 : {x : Ordinal} → def (od U) x → x o< & U - lemma2 {x} x<U = subst (λ k → k o< & U ) &iso (c<→o< (d→∋ U x<U)) - lemma1 : {x : Ordinal} → def (od U) x ∧ def (od (* x)) y → ¬ (& U o< y) - lemma1 {x} lt u<y = o<> u<y (ordtrans (lemma0 (proj2 lt)) (lemma2 (proj1 lt)) ) - lemma : ¬ ((& U) o< y ) → y o< osuc (& U) - lemma not with trio< y (& U) - lemma not | tri< a ¬b ¬c = ordtrans a <-osuc - lemma not | tri≈ ¬a refl ¬c = <-osuc - lemma not | tri> ¬a ¬b c = ⊥-elim (not c) -_∈_ : ( A B : HOD ) → Set n -A ∈ B = B ∋ A - -OPwr : (A : HOD ) → HOD -OPwr A = Ord ( sup-o (Ord (osuc (& A))) ( λ x A∋x → & ( A ∩ (* x)) ) ) - -Power : HOD → HOD -Power A = Replace (OPwr (Ord (& A))) ( λ x → A ∩ x ) --- {_} : ZFSet → ZFSet --- { x } = ( x , x ) -- better to use (x , x) directly - -union→ : (X z u : HOD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z -union→ X z u xx not = ⊥-elim ( not (& u) ( ⟪ proj1 xx - , subst ( λ k → odef k (& z)) (sym *iso) (proj2 xx) ⟫ )) -union← : (X z : HOD) (X∋z : Union X ∋ z) → ¬ ( (u : HOD ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) -union← X z UX∋z = FExists _ lemma UX∋z where - lemma : {y : Ordinal} → odef X y ∧ odef (* y) (& z) → ¬ ((u : HOD) → ¬ (X ∋ u) ∧ (u ∋ z)) - lemma {y} xx not = not (* y) ⟪ d→∋ X (proj1 xx) , proj2 xx ⟫ - -data infinite-d : ( x : Ordinal ) → Set n where - iφ : infinite-d o∅ - isuc : {x : Ordinal } → infinite-d x → - infinite-d (& ( Union (* x , (* x , * x ) ) )) - --- ω can be diverged in our case, since we have no restriction on the corresponding ordinal of a pair. --- We simply assumes infinite-d y has a maximum. --- --- This means that many of OD may not be HODs because of the & mapping divergence. --- We should have some axioms to prevent this such as & x o< next (odmax x). --- --- postulate --- ωmax : Ordinal --- <ωmax : {y : Ordinal} → infinite-d y → y o< ωmax --- --- infinite : HOD --- infinite = record { od = record { def = λ x → infinite-d x } ; odmax = ωmax ; <odmax = <ωmax } - -infinite : HOD -infinite = record { od = record { def = λ x → infinite-d x } ; odmax = next o∅ ; <odmax = lemma } where - u : (y : Ordinal ) → HOD - u y = Union (* y , (* y , * y)) - -- next< : {x y z : Ordinal} → x o< next z → y o< next x → y o< next z - lemma8 : {y : Ordinal} → & (* y , * y) o< next (odmax (* y , * y)) - lemma8 = ho< - --- (x,y) < next (omax x y) < next (osuc y) = next y - lemmaa : {x y : HOD} → & x o< & y → & (x , y) o< next (& y) - lemmaa {x} {y} x<y = subst (λ k → & (x , y) o< k ) (sym nexto≡) (subst (λ k → & (x , y) o< next k ) (sym (omax< _ _ x<y)) ho< ) - lemma81 : {y : Ordinal} → & (* y , * y) o< next (& (* y)) - lemma81 {y} = nexto=n (subst (λ k → & (* y , * y) o< k ) (cong (λ k → next k) (omxx _)) lemma8) - lemma9 : {y : Ordinal} → & (* y , (* y , * y)) o< next (& (* y , * y)) - lemma9 = lemmaa (c<→o< (case1 refl)) - lemma71 : {y : Ordinal} → & (* y , (* y , * y)) o< next (& (* y)) - lemma71 = next< lemma81 lemma9 - lemma1 : {y : Ordinal} → & (u y) o< next (osuc (& (* y , (* y , * y)))) - lemma1 = ho< - --- main recursion - lemma : {y : Ordinal} → infinite-d y → y o< next o∅ - lemma {o∅} iφ = x<nx - lemma (isuc {y} x) = next< (lemma x) (next< (subst (λ k → & (* y , (* y , * y)) o< next k) &iso lemma71 ) (nexto=n lemma1)) - -empty : (x : HOD ) → ¬ (od∅ ∋ x) -empty x = ¬x<0 - -_=h=_ : (x y : HOD) → Set n -x =h= y = od x == od y - -pair→ : ( x y t : HOD ) → (x , y) ∋ t → ( t =h= x ) ∨ ( t =h= y ) -pair→ x y t (case1 t≡x ) = case1 (subst₂ (λ j k → j =h= k ) *iso *iso (o≡→== t≡x )) -pair→ x y t (case2 t≡y ) = case2 (subst₂ (λ j k → j =h= k ) *iso *iso (o≡→== t≡y )) - -pair← : ( x y t : HOD ) → ( t =h= x ) ∨ ( t =h= y ) → (x , y) ∋ t -pair← x y t (case1 t=h=x) = case1 (cong (λ k → & k ) (==→o≡ t=h=x)) -pair← x y t (case2 t=h=y) = case2 (cong (λ k → & k ) (==→o≡ t=h=y)) - -o<→c< : {x y : Ordinal } → x o< y → (Ord x) ⊆ (Ord y) -o<→c< lt = record { incl = λ z → ordtrans z lt } - -⊆→o< : {x y : Ordinal } → (Ord x) ⊆ (Ord y) → x o< osuc y -⊆→o< {x} {y} lt with trio< x y -⊆→o< {x} {y} lt | tri< a ¬b ¬c = ordtrans a <-osuc -⊆→o< {x} {y} lt | tri≈ ¬a b ¬c = subst ( λ k → k o< osuc y) (sym b) <-osuc -⊆→o< {x} {y} lt | tri> ¬a ¬b c with (incl lt) (o<-subst c (sym &iso) refl ) -... | ttt = ⊥-elim ( o<¬≡ refl (o<-subst ttt &iso refl )) - -ψiso : {ψ : HOD → Set n} {x y : HOD } → ψ x → x ≡ y → ψ y -ψiso {ψ} t refl = t -selection : {ψ : HOD → Set n} {X y : HOD} → ((X ∋ y) ∧ ψ y) ⇔ (Select X ψ ∋ y) -selection {ψ} {X} {y} = ⟪ - ( λ cond → ⟪ proj1 cond , ψiso {ψ} (proj2 cond) (sym *iso) ⟫ ) - , ( λ select → ⟪ proj1 select , ψiso {ψ} (proj2 select) *iso ⟫ ) - ⟫ - -selection-in-domain : {ψ : HOD → Set n} {X y : HOD} → Select X ψ ∋ y → X ∋ y -selection-in-domain {ψ} {X} {y} lt = proj1 ((proj2 (selection {ψ} {X} )) lt) - -sup-c< : (ψ : HOD → HOD) → {X x : HOD} → X ∋ x → & (ψ x) o< (sup-o X (λ y X∋y → & (ψ (* y)))) -sup-c< ψ {X} {x} lt = subst (λ k → & (ψ k) o< _ ) *iso (sup-o< X lt ) -replacement← : {ψ : HOD → HOD} (X x : HOD) → X ∋ x → Replace X ψ ∋ ψ x -replacement← {ψ} X x lt = ⟪ sup-c< ψ {X} {x} lt , lemma ⟫ where - lemma : def (in-codomain X ψ) (& (ψ x)) - lemma not = ⊥-elim ( not ( & x ) ⟪ lt , cong (λ k → & (ψ k)) (sym *iso)⟫ ) -replacement→ : {ψ : HOD → HOD} (X x : HOD) → (lt : Replace X ψ ∋ x) → ¬ ( (y : HOD) → ¬ (x =h= ψ y)) -replacement→ {ψ} X x lt = contra-position lemma (lemma2 (proj2 lt)) where - lemma2 : ¬ ((y : Ordinal) → ¬ odef X y ∧ ((& x) ≡ & (ψ (* y)))) - → ¬ ((y : Ordinal) → ¬ odef X y ∧ (* (& x) =h= ψ (* y))) - lemma2 not not2 = not ( λ y d → not2 y ⟪ proj1 d , lemma3 (proj2 d)⟫) where - lemma3 : {y : Ordinal } → (& x ≡ & (ψ (* y))) → (* (& x) =h= ψ (* y)) - lemma3 {y} eq = subst (λ k → * (& x) =h= k ) *iso (o≡→== eq ) - lemma : ( (y : HOD) → ¬ (x =h= ψ y)) → ( (y : Ordinal) → ¬ odef X y ∧ (* (& x) =h= ψ (* y)) ) - lemma not y not2 = not (* y) (subst (λ k → k =h= ψ (* y)) *iso ( proj2 not2 )) - ---- ---- Power Set ---- ---- First consider ordinals in HOD ---- ---- A ∩ x = record { def = λ y → odef A y ∧ odef x y } subset of A --- --- -∩-≡ : { a b : HOD } → ({x : HOD } → (a ∋ x → b ∋ x)) → a =h= ( b ∩ a ) -∩-≡ {a} {b} inc = record { - eq→ = λ {x} x<a → ⟪ (subst (λ k → odef b k ) &iso (inc (d→∋ a x<a))) , x<a ⟫ ; - eq← = λ {x} x<a∩b → proj2 x<a∩b } --- --- Transitive Set case --- we have t ∋ x → Ord a ∋ x means t is a subset of Ord a, that is (Ord a) ∩ t =h= t --- OPwr (Ord a) is a sup of (Ord a) ∩ t, so OPwr (Ord a) ∋ t --- OPwr A = Ord ( sup-o ( λ x → & ( A ∩ (* x )) ) ) --- -ord-power← : (a : Ordinal ) (t : HOD) → ({x : HOD} → (t ∋ x → (Ord a) ∋ x)) → OPwr (Ord a) ∋ t -ord-power← a t t→A = subst (λ k → odef (OPwr (Ord a)) k ) (lemma1 lemma-eq) lemma where - lemma-eq : ((Ord a) ∩ t) =h= t - eq→ lemma-eq {z} w = proj2 w - eq← lemma-eq {z} w = ⟪ subst (λ k → odef (Ord a) k ) &iso ( t→A (d→∋ t w)) , w ⟫ - lemma1 : {a : Ordinal } { t : HOD } - → (eq : ((Ord a) ∩ t) =h= t) → & ((Ord a) ∩ (* (& t))) ≡ & t - lemma1 {a} {t} eq = subst (λ k → & ((Ord a) ∩ k) ≡ & t ) (sym *iso) (cong (λ k → & k ) (==→o≡ eq )) - lemma2 : (& t) o< (osuc (& (Ord a))) - lemma2 = ⊆→o≤ {t} {Ord a} (λ {x} x<t → subst (λ k → def (od (Ord a)) k) &iso (t→A (d→∋ t x<t))) - lemma : & ((Ord a) ∩ (* (& t)) ) o< sup-o (Ord (osuc (& (Ord a)))) (λ x lt → & ((Ord a) ∩ (* x))) - lemma = sup-o< _ lemma2 - --- --- Every set in HOD is a subset of Ordinals, so make OPwr (Ord (& A)) first --- then replace of all elements of the Power set by A ∩ y --- --- Power A = Replace (OPwr (Ord (& A))) ( λ y → A ∩ y ) - --- we have oly double negation form because of the replacement axiom --- -power→ : ( A t : HOD) → Power A ∋ t → {x : HOD} → t ∋ x → ¬ ¬ (A ∋ x) -power→ A t P∋t {x} t∋x = FExists _ lemma5 lemma4 where - a = & A - lemma2 : ¬ ( (y : HOD) → ¬ (t =h= (A ∩ y))) - lemma2 = replacement→ {λ x → A ∩ x} (OPwr (Ord (& A))) t P∋t - lemma3 : (y : HOD) → t =h= ( A ∩ y ) → ¬ ¬ (A ∋ x) - lemma3 y eq not = not (proj1 (eq→ eq t∋x)) - lemma4 : ¬ ((y : Ordinal) → ¬ (t =h= (A ∩ * y))) - lemma4 not = lemma2 ( λ y not1 → not (& y) (subst (λ k → t =h= ( A ∩ k )) (sym *iso) not1 )) - lemma5 : {y : Ordinal} → t =h= (A ∩ * y) → ¬ ¬ (odef A (& x)) - lemma5 {y} eq not = (lemma3 (* y) eq) not - -power← : (A t : HOD) → ({x : HOD} → (t ∋ x → A ∋ x)) → Power A ∋ t -power← A t t→A = ⟪ lemma1 , lemma2 ⟫ where - a = & A - lemma0 : {x : HOD} → t ∋ x → Ord a ∋ x - lemma0 {x} t∋x = c<→o< (t→A t∋x) - lemma3 : OPwr (Ord a) ∋ t - lemma3 = ord-power← a t lemma0 - lemma4 : (A ∩ * (& t)) ≡ t - lemma4 = let open ≡-Reasoning in begin - A ∩ * (& t) - ≡⟨ cong (λ k → A ∩ k) *iso ⟩ - A ∩ t - ≡⟨ sym (==→o≡ ( ∩-≡ {t} {A} t→A )) ⟩ - t - ∎ - sup1 : Ordinal - sup1 = sup-o (Ord (osuc (& (Ord (& A))))) (λ x A∋x → & ((Ord (& A)) ∩ (* x))) - lemma9 : def (od (Ord (Ordinals.osuc O (& (Ord (& A)))))) (& (Ord (& A))) - lemma9 = <-osuc - lemmab : & ((Ord (& A)) ∩ (* (& (Ord (& A) )))) o< sup1 - lemmab = sup-o< (Ord (osuc (& (Ord (& A))))) lemma9 - lemmad : Ord (osuc (& A)) ∋ t - lemmad = ⊆→o≤ (λ {x} lt → subst (λ k → def (od A) k ) &iso (t→A (d→∋ t lt))) - lemmac : ((Ord (& A)) ∩ (* (& (Ord (& A) )))) =h= Ord (& A) - lemmac = record { eq→ = lemmaf ; eq← = lemmag } where - lemmaf : {x : Ordinal} → def (od ((Ord (& A)) ∩ (* (& (Ord (& A)))))) x → def (od (Ord (& A))) x - lemmaf {x} lt = proj1 lt - lemmag : {x : Ordinal} → def (od (Ord (& A))) x → def (od ((Ord (& A)) ∩ (* (& (Ord (& A)))))) x - lemmag {x} lt = ⟪ lt , subst (λ k → def (od k) x) (sym *iso) lt ⟫ - lemmae : & ((Ord (& A)) ∩ (* (& (Ord (& A))))) ≡ & (Ord (& A)) - lemmae = cong (λ k → & k ) ( ==→o≡ lemmac) - lemma7 : def (od (OPwr (Ord (& A)))) (& t) - lemma7 with osuc-≡< lemmad - lemma7 | case2 lt = ordtrans (c<→o< lt) (subst (λ k → k o< sup1) lemmae lemmab ) - lemma7 | case1 eq with osuc-≡< (⊆→o≤ {* (& t)} {* (& (Ord (& t)))} (λ {x} lt → lemmah lt )) where - lemmah : {x : Ordinal } → def (od (* (& t))) x → def (od (* (& (Ord (& t))))) x - lemmah {x} lt = subst (λ k → def (od k) x ) (sym *iso) (subst (λ k → k o< (& t)) - &iso - (c<→o< (subst₂ (λ j k → def (od j) k) *iso (sym &iso) lt ))) - lemma7 | case1 eq | case1 eq1 = subst (λ k → k o< sup1) (trans lemmae lemmai) lemmab where - lemmai : & (Ord (& A)) ≡ & t - lemmai = let open ≡-Reasoning in begin - & (Ord (& A)) - ≡⟨ sym (cong (λ k → & (Ord k)) eq) ⟩ - & (Ord (& t)) - ≡⟨ sym &iso ⟩ - & (* (& (Ord (& t)))) - ≡⟨ sym eq1 ⟩ - & (* (& t)) - ≡⟨ &iso ⟩ - & t - ∎ - lemma7 | case1 eq | case2 lt = ordtrans lemmaj (subst (λ k → k o< sup1) lemmae lemmab ) where - lemmak : & (* (& (Ord (& t)))) ≡ & (Ord (& A)) - lemmak = let open ≡-Reasoning in begin - & (* (& (Ord (& t)))) - ≡⟨ &iso ⟩ - & (Ord (& t)) - ≡⟨ cong (λ k → & (Ord k)) eq ⟩ - & (Ord (& A)) - ∎ - lemmaj : & t o< & (Ord (& A)) - lemmaj = subst₂ (λ j k → j o< k ) &iso lemmak lt - lemma1 : & t o< sup-o (OPwr (Ord (& A))) (λ x lt → & (A ∩ (* x))) - lemma1 = subst (λ k → & k o< sup-o (OPwr (Ord (& A))) (λ x lt → & (A ∩ (* x)))) - lemma4 (sup-o< (OPwr (Ord (& A))) lemma7 ) - lemma2 : def (in-codomain (OPwr (Ord (& A))) (_∩_ A)) (& t) - lemma2 not = ⊥-elim ( not (& t) ⟪ lemma3 , lemma6 ⟫ ) where - lemma6 : & t ≡ & (A ∩ * (& t)) - lemma6 = cong ( λ k → & k ) (==→o≡ (subst (λ k → t =h= (A ∩ k)) (sym *iso) ( ∩-≡ {t} {A} t→A ))) - - -extensionality0 : {A B : HOD } → ((z : HOD) → (A ∋ z) ⇔ (B ∋ z)) → A =h= B -eq→ (extensionality0 {A} {B} eq ) {x} d = odef-iso {A} {B} (sym &iso) (proj1 (eq (* x))) d -eq← (extensionality0 {A} {B} eq ) {x} d = odef-iso {B} {A} (sym &iso) (proj2 (eq (* x))) d - -extensionality : {A B w : HOD } → ((z : HOD ) → (A ∋ z) ⇔ (B ∋ z)) → (w ∋ A) ⇔ (w ∋ B) -proj1 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) ( ==→o≡ (extensionality0 {A} {B} eq) ) d -proj2 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) (sym ( ==→o≡ (extensionality0 {A} {B} eq) )) d - -infinity∅ : infinite ∋ od∅ -infinity∅ = subst (λ k → odef infinite k ) lemma iφ where - lemma : o∅ ≡ & od∅ - lemma = let open ≡-Reasoning in begin - o∅ - ≡⟨ sym &iso ⟩ - & ( * o∅ ) - ≡⟨ cong ( λ k → & k ) o∅≡od∅ ⟩ - & od∅ - ∎ -infinity : (x : HOD) → infinite ∋ x → infinite ∋ Union (x , (x , x )) -infinity x lt = subst (λ k → odef infinite k ) lemma (isuc {& x} lt) where - lemma : & (Union (* (& x) , (* (& x) , * (& x)))) - ≡ & (Union (x , (x , x))) - lemma = cong (λ k → & (Union ( k , ( k , k ) ))) *iso - -isZF : IsZF (HOD ) _∋_ _=h=_ od∅ _,_ Union Power Select Replace infinite -isZF = record { - isEquivalence = record { refl = ==-refl ; sym = ==-sym; trans = ==-trans } - ; pair→ = pair→ - ; pair← = pair← - ; union→ = union→ - ; union← = union← - ; empty = empty - ; power→ = power→ - ; power← = power← - ; extensionality = λ {A} {B} {w} → extensionality {A} {B} {w} - ; ε-induction = ε-induction - ; infinity∅ = infinity∅ - ; infinity = infinity - ; selection = λ {X} {ψ} {y} → selection {X} {ψ} {y} - ; replacement← = replacement← - ; replacement→ = λ {ψ} → replacement→ {ψ} - } - -HOD→ZF : ZF -HOD→ZF = record { - ZFSet = HOD - ; _∋_ = _∋_ - ; _≈_ = _=h=_ - ; ∅ = od∅ - ; _,_ = _,_ - ; Union = Union - ; Power = Power - ; Select = Select - ; Replace = Replace - ; infinite = infinite - ; isZF = isZF - } - -
--- a/ODC.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,128 +0,0 @@ -{-# OPTIONS --allow-unsolved-metas #-} -open import Level -open import Ordinals -module ODC {n : Level } (O : Ordinals {n} ) where - -open import zf -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -open import Relation.Binary.PropositionalEquality -open import Data.Nat.Properties -open import Data.Empty -open import Relation.Nullary -open import Relation.Binary -open import Relation.Binary.Core - -import OrdUtil -open import logic -open import nat -import OD -import ODUtil - -open inOrdinal O -open OD O -open OD.OD -open OD._==_ -open ODAxiom odAxiom -open ODUtil O - -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -open OrdUtil O - - -open HOD - -open _∧_ - -postulate - -- mimimul and x∋minimal is an Axiom of choice - minimal : (x : HOD ) → ¬ (x =h= od∅ )→ HOD - -- this should be ¬ (x =h= od∅ )→ ∃ ox → x ∋ Ord ox ( minimum of x ) - x∋minimal : (x : HOD ) → ( ne : ¬ (x =h= od∅ ) ) → odef x ( & ( minimal x ne ) ) - -- minimality (may proved by ε-induction with LEM) - minimal-1 : (x : HOD ) → ( ne : ¬ (x =h= od∅ ) ) → (y : HOD ) → ¬ ( odef (minimal x ne) (& y)) ∧ (odef x (& y) ) - - --- --- Axiom of choice in intutionistic logic implies the exclude middle --- https://plato.stanford.edu/entries/axiom-choice/#AxiChoLog --- - -pred-od : ( p : Set n ) → HOD -pred-od p = record { od = record { def = λ x → (x ≡ o∅) ∧ p } ; - odmax = osuc o∅; <odmax = λ x → subst (λ k → k o< osuc o∅) (sym (proj1 x)) <-osuc } - -ppp : { p : Set n } { a : HOD } → pred-od p ∋ a → p -ppp {p} {a} d = proj2 d - -p∨¬p : ( p : Set n ) → p ∨ ( ¬ p ) -- assuming axiom of choice -p∨¬p p with is-o∅ ( & (pred-od p )) -p∨¬p p | yes eq = case2 (¬p eq) where - ps = pred-od p - eqo∅ : ps =h= od∅ → & ps ≡ o∅ - eqo∅ eq = subst (λ k → & ps ≡ k) ord-od∅ ( cong (λ k → & k ) (==→o≡ eq)) - lemma : ps =h= od∅ → p → ⊥ - lemma eq p0 = ¬x<0 {& ps} (eq→ eq record { proj1 = eqo∅ eq ; proj2 = p0 } ) - ¬p : (& ps ≡ o∅) → p → ⊥ - ¬p eq = lemma ( subst₂ (λ j k → j =h= k ) *iso o∅≡od∅ ( o≡→== eq )) -p∨¬p p | no ¬p = case1 (ppp {p} {minimal ps (λ eq → ¬p (eqo∅ eq))} lemma) where - ps = pred-od p - eqo∅ : ps =h= od∅ → & ps ≡ o∅ - eqo∅ eq = subst (λ k → & ps ≡ k) ord-od∅ ( cong (λ k → & k ) (==→o≡ eq)) - lemma : ps ∋ minimal ps (λ eq → ¬p (eqo∅ eq)) - lemma = x∋minimal ps (λ eq → ¬p (eqo∅ eq)) - -decp : ( p : Set n ) → Dec p -- assuming axiom of choice -decp p with p∨¬p p -decp p | case1 x = yes x -decp p | case2 x = no x - -∋-p : (A x : HOD ) → Dec ( A ∋ x ) -∋-p A x with p∨¬p ( A ∋ x ) -- LEM -∋-p A x | case1 t = yes t -∋-p A x | case2 t = no (λ x → t x) - -double-neg-eilm : {A : Set n} → ¬ ¬ A → A -- we don't have this in intutionistic logic -double-neg-eilm {A} notnot with decp A -- assuming axiom of choice -... | yes p = p -... | no ¬p = ⊥-elim ( notnot ¬p ) - -open _⊆_ - -power→⊆ : ( A t : HOD) → Power A ∋ t → t ⊆ A -power→⊆ A t PA∋t = record { incl = λ {x} t∋x → double-neg-eilm (t1 t∋x) } where - t1 : {x : HOD } → t ∋ x → ¬ ¬ (A ∋ x) - t1 = power→ A t PA∋t - -OrdP : ( x : Ordinal ) ( y : HOD ) → Dec ( Ord x ∋ y ) -OrdP x y with trio< x (& y) -OrdP x y | tri< a ¬b ¬c = no ¬c -OrdP x y | tri≈ ¬a refl ¬c = no ( o<¬≡ refl ) -OrdP x y | tri> ¬a ¬b c = yes c - -open import zfc - -HOD→ZFC : ZFC -HOD→ZFC = record { - ZFSet = HOD - ; _∋_ = _∋_ - ; _≈_ = _=h=_ - ; ∅ = od∅ - ; Select = Select - ; isZFC = isZFC - } where - -- infixr 200 _∈_ - -- infixr 230 _∩_ _∪_ - isZFC : IsZFC (HOD ) _∋_ _=h=_ od∅ Select - isZFC = record { - choice-func = choice-func ; - choice = choice - } where - -- Axiom of choice ( is equivalent to the existence of minimal in our case ) - -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] - choice-func : (X : HOD ) → {x : HOD } → ¬ ( x =h= od∅ ) → ( X ∋ x ) → HOD - choice-func X {x} not X∋x = minimal x not - choice : (X : HOD ) → {A : HOD } → ( X∋A : X ∋ A ) → (not : ¬ ( A =h= od∅ )) → A ∋ choice-func X not X∋A - choice X {A} X∋A not = x∋minimal A not -
--- a/ODUtil.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,179 +0,0 @@ -{-# OPTIONS --allow-unsolved-metas #-} -open import Level -open import Ordinals -module ODUtil {n : Level } (O : Ordinals {n} ) where - -open import zf -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -open import Relation.Binary.PropositionalEquality hiding ( [_] ) -open import Data.Nat.Properties -open import Data.Empty -open import Relation.Nullary -open import Relation.Binary hiding ( _⇔_ ) - -open import logic -open import nat - -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -import OrdUtil -open OrdUtil O - -import OD -open OD O -open OD.OD -open ODAxiom odAxiom - -open HOD -open _⊆_ -open _∧_ -open _==_ - -cseq : HOD → HOD -cseq x = record { od = record { def = λ y → odef x (osuc y) } ; odmax = osuc (odmax x) ; <odmax = lemma } where - lemma : {y : Ordinal} → def (od x) (osuc y) → y o< osuc (odmax x) - lemma {y} lt = ordtrans <-osuc (ordtrans (<odmax x lt) <-osuc ) - - -pair-xx<xy : {x y : HOD} → & (x , x) o< osuc (& (x , y) ) -pair-xx<xy {x} {y} = ⊆→o≤ lemma where - lemma : {z : Ordinal} → def (od (x , x)) z → def (od (x , y)) z - lemma {z} (case1 refl) = case1 refl - lemma {z} (case2 refl) = case1 refl - -pair-<xy : {x y : HOD} → {n : Ordinal} → & x o< next n → & y o< next n → & (x , y) o< next n -pair-<xy {x} {y} {o} x<nn y<nn with trio< (& x) (& y) | inspect (omax (& x)) (& y) -... | tri< a ¬b ¬c | record { eq = eq1 } = next< (subst (λ k → k o< next o ) (sym eq1) (osuc<nx y<nn)) ho< -... | tri> ¬a ¬b c | record { eq = eq1 } = next< (subst (λ k → k o< next o ) (sym eq1) (osuc<nx x<nn)) ho< -... | tri≈ ¬a b ¬c | record { eq = eq1 } = next< (subst (λ k → k o< next o ) (omax≡ _ _ b) (subst (λ k → osuc k o< next o) b (osuc<nx x<nn))) ho< - --- another form of infinite --- pair-ord< : {x : Ordinal } → Set n -pair-ord< : {x : HOD } → ( {y : HOD } → & y o< next (odmax y) ) → & ( x , x ) o< next (& x) -pair-ord< {x} ho< = subst (λ k → & (x , x) o< k ) lemmab0 lemmab1 where - lemmab0 : next (odmax (x , x)) ≡ next (& x) - lemmab0 = trans (cong (λ k → next k) (omxx _)) (sym nexto≡) - lemmab1 : & (x , x) o< next ( odmax (x , x)) - lemmab1 = ho< - -trans-⊆ : { A B C : HOD} → A ⊆ B → B ⊆ C → A ⊆ C -trans-⊆ A⊆B B⊆C = record { incl = λ x → incl B⊆C (incl A⊆B x) } - -refl-⊆ : {A : HOD} → A ⊆ A -refl-⊆ {A} = record { incl = λ x → x } - -od⊆→o≤ : {x y : HOD } → x ⊆ y → & x o< osuc (& y) -od⊆→o≤ {x} {y} lt = ⊆→o≤ {x} {y} (λ {z} x>z → subst (λ k → def (od y) k ) &iso (incl lt (d→∋ x x>z))) - -subset-lemma : {A x : HOD } → ( {y : HOD } → x ∋ y → (A ∩ x ) ∋ y ) ⇔ ( x ⊆ A ) -subset-lemma {A} {x} = record { - proj1 = λ lt → record { incl = λ x∋z → proj1 (lt x∋z) } - ; proj2 = λ x⊆A lt → ⟪ incl x⊆A lt , lt ⟫ - } - - -ω<next-o∅ : {y : Ordinal} → infinite-d y → y o< next o∅ -ω<next-o∅ {y} lt = <odmax infinite lt - -nat→ω : Nat → HOD -nat→ω Zero = od∅ -nat→ω (Suc y) = Union (nat→ω y , (nat→ω y , nat→ω y)) - -ω→nato : {y : Ordinal} → infinite-d y → Nat -ω→nato iφ = Zero -ω→nato (isuc lt) = Suc (ω→nato lt) - -ω→nat : (n : HOD) → infinite ∋ n → Nat -ω→nat n = ω→nato - -ω∋nat→ω : {n : Nat} → def (od infinite) (& (nat→ω n)) -ω∋nat→ω {Zero} = subst (λ k → def (od infinite) k) (sym ord-od∅) iφ -ω∋nat→ω {Suc n} = subst (λ k → def (od infinite) k) lemma (isuc ( ω∋nat→ω {n})) where - lemma : & (Union (* (& (nat→ω n)) , (* (& (nat→ω n)) , * (& (nat→ω n))))) ≡ & (nat→ω (Suc n)) - lemma = subst (λ k → & (Union (k , ( k , k ))) ≡ & (nat→ω (Suc n))) (sym *iso) refl - -pair1 : { x y : HOD } → (x , y ) ∋ x -pair1 = case1 refl - -pair2 : { x y : HOD } → (x , y ) ∋ y -pair2 = case2 refl - -single : {x y : HOD } → (x , x ) ∋ y → x ≡ y -single (case1 eq) = ==→o≡ ( ord→== (sym eq) ) -single (case2 eq) = ==→o≡ ( ord→== (sym eq) ) - -open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) --- postulate f-extensionality : { n m : Level} → HE.Extensionality n m - -ω-prev-eq1 : {x y : Ordinal} → & (Union (* y , (* y , * y))) ≡ & (Union (* x , (* x , * x))) → ¬ (x o< y) -ω-prev-eq1 {x} {y} eq x<y = eq→ (ord→== eq) {& (* y)} (λ not2 → not2 (& (* y , * y)) - ⟪ case2 refl , subst (λ k → odef k (& (* y))) (sym *iso) (case1 refl)⟫ ) (λ u → lemma u ) where - lemma : (u : Ordinal) → ¬ def (od (* x , (* x , * x))) u ∧ def (od (* u)) (& (* y)) - lemma u t with proj1 t - lemma u t | case1 u=x = o<> (c<→o< {* y} {* u} (proj2 t)) (subst₂ (λ j k → j o< k ) - (trans (sym &iso) (trans (sym u=x) (sym &iso)) ) (sym &iso) x<y ) -- x ≡ & (* u) - lemma u t | case2 u=xx = o<¬≡ (lemma1 (subst (λ k → odef k (& (* y)) ) (trans (cong (λ k → * k ) u=xx) *iso ) (proj2 t))) x<y where - lemma1 : {x y : Ordinal } → (* x , * x ) ∋ * y → x ≡ y -- y = x ∈ ( x , x ) = u - lemma1 (case1 eq) = subst₂ (λ j k → j ≡ k ) &iso &iso (sym eq) - lemma1 (case2 eq) = subst₂ (λ j k → j ≡ k ) &iso &iso (sym eq) - -ω-prev-eq : {x y : Ordinal} → & (Union (* y , (* y , * y))) ≡ & (Union (* x , (* x , * x))) → x ≡ y -ω-prev-eq {x} {y} eq with trio< x y -ω-prev-eq {x} {y} eq | tri< a ¬b ¬c = ⊥-elim (ω-prev-eq1 eq a) -ω-prev-eq {x} {y} eq | tri≈ ¬a b ¬c = b -ω-prev-eq {x} {y} eq | tri> ¬a ¬b c = ⊥-elim (ω-prev-eq1 (sym eq) c) - -ω-∈s : (x : HOD) → Union ( x , (x , x)) ∋ x -ω-∈s x not = not (& (x , x)) ⟪ case2 refl , subst (λ k → odef k (& x) ) (sym *iso) (case1 refl) ⟫ - -ωs≠0 : (x : HOD) → ¬ ( Union ( x , (x , x)) ≡ od∅ ) -ωs≠0 y eq = ⊥-elim ( ¬x<0 (subst (λ k → & y o< k ) ord-od∅ (c<→o< (subst (λ k → odef k (& y )) eq (ω-∈s y) ))) ) - -nat→ω-iso : {i : HOD} → (lt : infinite ∋ i ) → nat→ω ( ω→nat i lt ) ≡ i -nat→ω-iso {i} = ε-induction {λ i → (lt : infinite ∋ i ) → nat→ω ( ω→nat i lt ) ≡ i } ind i where - ind : {x : HOD} → ({y : HOD} → x ∋ y → (lt : infinite ∋ y) → nat→ω (ω→nat y lt) ≡ y) → - (lt : infinite ∋ x) → nat→ω (ω→nat x lt) ≡ x - ind {x} prev lt = ind1 lt *iso where - ind1 : {ox : Ordinal } → (ltd : infinite-d ox ) → * ox ≡ x → nat→ω (ω→nato ltd) ≡ x - ind1 {o∅} iφ refl = sym o∅≡od∅ - ind1 (isuc {x₁} ltd) ox=x = begin - nat→ω (ω→nato (isuc ltd) ) - ≡⟨⟩ - Union (nat→ω (ω→nato ltd) , (nat→ω (ω→nato ltd) , nat→ω (ω→nato ltd))) - ≡⟨ cong (λ k → Union (k , (k , k ))) lemma ⟩ - Union (* x₁ , (* x₁ , * x₁)) - ≡⟨ trans ( sym *iso) ox=x ⟩ - x - ∎ where - open ≡-Reasoning - lemma0 : x ∋ * x₁ - lemma0 = subst (λ k → odef k (& (* x₁))) (trans (sym *iso) ox=x) (λ not → not - (& (* x₁ , * x₁)) ⟪ pair2 , subst (λ k → odef k (& (* x₁))) (sym *iso) pair1 ⟫ ) - lemma1 : infinite ∋ * x₁ - lemma1 = subst (λ k → odef infinite k) (sym &iso) ltd - lemma3 : {x y : Ordinal} → (ltd : infinite-d x ) (ltd1 : infinite-d y ) → y ≡ x → ltd ≅ ltd1 - lemma3 iφ iφ refl = HE.refl - lemma3 iφ (isuc {y} ltd1) eq = ⊥-elim ( ¬x<0 (subst₂ (λ j k → j o< k ) &iso eq (c<→o< (ω-∈s (* y)) ))) - lemma3 (isuc {y} ltd) iφ eq = ⊥-elim ( ¬x<0 (subst₂ (λ j k → j o< k ) &iso (sym eq) (c<→o< (ω-∈s (* y)) ))) - lemma3 (isuc {x} ltd) (isuc {y} ltd1) eq with lemma3 ltd ltd1 (ω-prev-eq (sym eq)) - ... | t = HE.cong₂ (λ j k → isuc {j} k ) (HE.≡-to-≅ (ω-prev-eq eq)) t - lemma2 : {x y : Ordinal} → (ltd : infinite-d x ) (ltd1 : infinite-d y ) → y ≡ x → ω→nato ltd ≡ ω→nato ltd1 - lemma2 {x} {y} ltd ltd1 eq = lemma6 eq (lemma3 {x} {y} ltd ltd1 eq) where - lemma6 : {x y : Ordinal} → {ltd : infinite-d x } {ltd1 : infinite-d y } → y ≡ x → ltd ≅ ltd1 → ω→nato ltd ≡ ω→nato ltd1 - lemma6 refl HE.refl = refl - lemma : nat→ω (ω→nato ltd) ≡ * x₁ - lemma = trans (cong (λ k → nat→ω k) (lemma2 {x₁} {_} ltd (subst (λ k → infinite-d k ) (sym &iso) ltd) &iso ) ) ( prev {* x₁} lemma0 lemma1 ) - -ω→nat-iso : {i : Nat} → ω→nat ( nat→ω i ) (ω∋nat→ω {i}) ≡ i -ω→nat-iso {i} = lemma i (ω∋nat→ω {i}) *iso where - lemma : {x : Ordinal } → ( i : Nat ) → (ltd : infinite-d x ) → * x ≡ nat→ω i → ω→nato ltd ≡ i - lemma {x} Zero iφ eq = refl - lemma {x} (Suc i) iφ eq = ⊥-elim ( ωs≠0 (nat→ω i) (trans (sym eq) o∅≡od∅ )) -- Union (nat→ω i , (nat→ω i , nat→ω i)) ≡ od∅ - lemma Zero (isuc {x} ltd) eq = ⊥-elim ( ωs≠0 (* x) (subst (λ k → k ≡ od∅ ) *iso eq )) - lemma (Suc i) (isuc {x} ltd) eq = cong (λ k → Suc k ) (lemma i ltd (lemma1 eq) ) where -- * x ≡ nat→ω i - lemma1 : * (& (Union (* x , (* x , * x)))) ≡ Union (nat→ω i , (nat→ω i , nat→ω i)) → * x ≡ nat→ω i - lemma1 eq = subst (λ k → * x ≡ k ) *iso (cong (λ k → * k) - ( ω-prev-eq (subst (λ k → _ ≡ k ) &iso (cong (λ k → & k ) (sym - (subst (λ k → _ ≡ Union ( k , ( k , k ))) (sym *iso ) eq )))))) -
--- a/OPair.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,213 +0,0 @@ -{-# OPTIONS --allow-unsolved-metas #-} - -open import Level -open import Ordinals -module OPair {n : Level } (O : Ordinals {n}) where - -open import zf -open import logic -import OD -import ODUtil -import OrdUtil - -open import Relation.Nullary -open import Relation.Binary -open import Data.Empty -open import Relation.Binary -open import Relation.Binary.Core -open import Relation.Binary.PropositionalEquality -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) - -open OD O -open OD.OD -open OD.HOD -open ODAxiom odAxiom - -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -open OrdUtil O -open ODUtil O - -open _∧_ -open _∨_ -open Bool - -open _==_ - -<_,_> : (x y : HOD) → HOD -< x , y > = (x , x ) , (x , y ) - -exg-pair : { x y : HOD } → (x , y ) =h= ( y , x ) -exg-pair {x} {y} = record { eq→ = left ; eq← = right } where - left : {z : Ordinal} → odef (x , y) z → odef (y , x) z - left (case1 t) = case2 t - left (case2 t) = case1 t - right : {z : Ordinal} → odef (y , x) z → odef (x , y) z - right (case1 t) = case2 t - right (case2 t) = case1 t - -ord≡→≡ : { x y : HOD } → & x ≡ & y → x ≡ y -ord≡→≡ eq = subst₂ (λ j k → j ≡ k ) *iso *iso ( cong ( λ k → * k ) eq ) - -od≡→≡ : { x y : Ordinal } → * x ≡ * y → x ≡ y -od≡→≡ eq = subst₂ (λ j k → j ≡ k ) &iso &iso ( cong ( λ k → & k ) eq ) - -eq-prod : { x x' y y' : HOD } → x ≡ x' → y ≡ y' → < x , y > ≡ < x' , y' > -eq-prod refl refl = refl - -xx=zy→x=y : {x y z : HOD } → ( x , x ) =h= ( z , y ) → x ≡ y -xx=zy→x=y {x} {y} eq with trio< (& x) (& y) -xx=zy→x=y {x} {y} eq | tri< a ¬b ¬c with eq← eq {& y} (case2 refl) -xx=zy→x=y {x} {y} eq | tri< a ¬b ¬c | case1 s = ⊥-elim ( o<¬≡ (sym s) a ) -xx=zy→x=y {x} {y} eq | tri< a ¬b ¬c | case2 s = ⊥-elim ( o<¬≡ (sym s) a ) -xx=zy→x=y {x} {y} eq | tri≈ ¬a b ¬c = ord≡→≡ b -xx=zy→x=y {x} {y} eq | tri> ¬a ¬b c with eq← eq {& y} (case2 refl) -xx=zy→x=y {x} {y} eq | tri> ¬a ¬b c | case1 s = ⊥-elim ( o<¬≡ s c ) -xx=zy→x=y {x} {y} eq | tri> ¬a ¬b c | case2 s = ⊥-elim ( o<¬≡ s c ) - -prod-eq : { x x' y y' : HOD } → < x , y > =h= < x' , y' > → (x ≡ x' ) ∧ ( y ≡ y' ) -prod-eq {x} {x'} {y} {y'} eq = ⟪ lemmax , lemmay ⟫ where - lemma2 : {x y z : HOD } → ( x , x ) =h= ( z , y ) → z ≡ y - lemma2 {x} {y} {z} eq = trans (sym (xx=zy→x=y lemma3 )) ( xx=zy→x=y eq ) where - lemma3 : ( x , x ) =h= ( y , z ) - lemma3 = ==-trans eq exg-pair - lemma1 : {x y : HOD } → ( x , x ) =h= ( y , y ) → x ≡ y - lemma1 {x} {y} eq with eq← eq {& y} (case2 refl) - lemma1 {x} {y} eq | case1 s = ord≡→≡ (sym s) - lemma1 {x} {y} eq | case2 s = ord≡→≡ (sym s) - lemma4 : {x y z : HOD } → ( x , y ) =h= ( x , z ) → y ≡ z - lemma4 {x} {y} {z} eq with eq← eq {& z} (case2 refl) - lemma4 {x} {y} {z} eq | case1 s with ord≡→≡ s -- x ≡ z - ... | refl with lemma2 (==-sym eq ) - ... | refl = refl - lemma4 {x} {y} {z} eq | case2 s = ord≡→≡ (sym s) -- y ≡ z - lemmax : x ≡ x' - lemmax with eq→ eq {& (x , x)} (case1 refl) - lemmax | case1 s = lemma1 (ord→== s ) -- (x,x)≡(x',x') - lemmax | case2 s with lemma2 (ord→== s ) -- (x,x)≡(x',y') with x'≡y' - ... | refl = lemma1 (ord→== s ) - lemmay : y ≡ y' - lemmay with lemmax - ... | refl with lemma4 eq -- with (x,y)≡(x,y') - ... | eq1 = lemma4 (ord→== (cong (λ k → & k ) eq1 )) - --- --- unlike ordered pair, ZFProduct is not a HOD - -data ord-pair : (p : Ordinal) → Set n where - pair : (x y : Ordinal ) → ord-pair ( & ( < * x , * y > ) ) - -ZFProduct : OD -ZFProduct = record { def = λ x → ord-pair x } - --- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) --- eq-pair : { x x' y y' : Ordinal } → x ≡ x' → y ≡ y' → pair x y ≅ pair x' y' --- eq-pair refl refl = HE.refl - -pi1 : { p : Ordinal } → ord-pair p → Ordinal -pi1 ( pair x y) = x - -π1 : { p : HOD } → def ZFProduct (& p) → HOD -π1 lt = * (pi1 lt ) - -pi2 : { p : Ordinal } → ord-pair p → Ordinal -pi2 ( pair x y ) = y - -π2 : { p : HOD } → def ZFProduct (& p) → HOD -π2 lt = * (pi2 lt ) - -op-cons : { ox oy : Ordinal } → def ZFProduct (& ( < * ox , * oy > )) -op-cons {ox} {oy} = pair ox oy - -def-subst : {Z : OD } {X : Ordinal }{z : OD } {x : Ordinal }→ def Z X → Z ≡ z → X ≡ x → def z x -def-subst df refl refl = df - -p-cons : ( x y : HOD ) → def ZFProduct (& ( < x , y >)) -p-cons x y = def-subst {_} {_} {ZFProduct} {& (< x , y >)} (pair (& x) ( & y )) refl ( - let open ≡-Reasoning in begin - & < * (& x) , * (& y) > - ≡⟨ cong₂ (λ j k → & < j , k >) *iso *iso ⟩ - & < x , y > - ∎ ) - -op-iso : { op : Ordinal } → (q : ord-pair op ) → & < * (pi1 q) , * (pi2 q) > ≡ op -op-iso (pair ox oy) = refl - -p-iso : { x : HOD } → (p : def ZFProduct (& x) ) → < π1 p , π2 p > ≡ x -p-iso {x} p = ord≡→≡ (op-iso p) - -p-pi1 : { x y : HOD } → (p : def ZFProduct (& < x , y >) ) → π1 p ≡ x -p-pi1 {x} {y} p = proj1 ( prod-eq ( ord→== (op-iso p) )) - -p-pi2 : { x y : HOD } → (p : def ZFProduct (& < x , y >) ) → π2 p ≡ y -p-pi2 {x} {y} p = proj2 ( prod-eq ( ord→== (op-iso p))) - -ω-pair : {x y : HOD} → {m : Ordinal} → & x o< next m → & y o< next m → & (x , y) o< next m -ω-pair lx ly = next< (omax<nx lx ly ) ho< - -ω-opair : {x y : HOD} → {m : Ordinal} → & x o< next m → & y o< next m → & < x , y > o< next m -ω-opair {x} {y} {m} lx ly = lemma0 where - lemma0 : & < x , y > o< next m - lemma0 = osucprev (begin - osuc (& < x , y >) - <⟨ osuc<nx ho< ⟩ - next (omax (& (x , x)) (& (x , y))) - ≡⟨ cong (λ k → next k) (sym ( omax≤ _ _ pair-xx<xy )) ⟩ - next (osuc (& (x , y))) - ≡⟨ sym (nexto≡) ⟩ - next (& (x , y)) - ≤⟨ x<ny→≤next (ω-pair lx ly) ⟩ - next m - ∎ ) where - open o≤-Reasoning O - -_⊗_ : (A B : HOD) → HOD -A ⊗ B = Union ( Replace B (λ b → Replace A (λ a → < a , b > ) )) - -product→ : {A B a b : HOD} → A ∋ a → B ∋ b → ( A ⊗ B ) ∋ < a , b > -product→ {A} {B} {a} {b} A∋a B∋b = λ t → t (& (Replace A (λ a → < a , b >))) - ⟪ lemma1 , subst (λ k → odef k (& < a , b >)) (sym *iso) lemma2 ⟫ where - lemma1 : odef (Replace B (λ b₁ → Replace A (λ a₁ → < a₁ , b₁ >))) (& (Replace A (λ a₁ → < a₁ , b >))) - lemma1 = replacement← B b B∋b - lemma2 : odef (Replace A (λ a₁ → < a₁ , b >)) (& < a , b >) - lemma2 = replacement← A a A∋a - -x<nextA : {A x : HOD} → A ∋ x → & x o< next (odmax A) -x<nextA {A} {x} A∋x = ordtrans (c<→o< {x} {A} A∋x) ho< - -A<Bnext : {A B x : HOD} → & A o< & B → A ∋ x → & x o< next (odmax B) -A<Bnext {A} {B} {x} lt A∋x = osucprev (begin - osuc (& x) - <⟨ osucc (c<→o< A∋x) ⟩ - osuc (& A) - <⟨ osucc lt ⟩ - osuc (& B) - <⟨ osuc<nx ho< ⟩ - next (odmax B) - ∎ ) where open o≤-Reasoning O - -ZFP : (A B : HOD) → HOD -ZFP A B = record { od = record { def = λ x → ord-pair x ∧ ((p : ord-pair x ) → odef A (pi1 p) ∧ odef B (pi2 p) )} ; - odmax = omax (next (odmax A)) (next (odmax B)) ; <odmax = λ {y} px → lemma y px } - where - lemma : (y : Ordinal) → ( ord-pair y ∧ ((p : ord-pair y) → odef A (pi1 p) ∧ odef B (pi2 p))) → y o< omax (next (odmax A)) (next (odmax B)) - lemma y lt with proj1 lt - lemma p lt | pair x y with trio< (& A) (& B) - lemma p lt | pair x y | tri< a ¬b ¬c = ordtrans (ω-opair (A<Bnext a (subst (λ k → odef A k ) (sym &iso) - (proj1 (proj2 lt (pair x y))))) (lemma1 (proj2 (proj2 lt (pair x y))))) (omax-y _ _ ) where - lemma1 : odef B y → & (* y) o< next (HOD.odmax B) - lemma1 lt = x<nextA {B} (d→∋ B lt) - lemma p lt | pair x y | tri≈ ¬a b ¬c = ordtrans (ω-opair (x<nextA {A} (d→∋ A ((proj1 (proj2 lt (pair x y)))))) lemma2 ) (omax-x _ _ ) where - lemma2 : & (* y) o< next (HOD.odmax A) - lemma2 = ordtrans ( subst (λ k → & (* y) o< k ) (sym b) (c<→o< (d→∋ B ((proj2 (proj2 lt (pair x y))))))) ho< - lemma p lt | pair x y | tri> ¬a ¬b c = ordtrans (ω-opair (x<nextA {A} (d→∋ A ((proj1 (proj2 lt (pair x y)))))) - (A<Bnext c (subst (λ k → odef B k ) (sym &iso) (proj2 (proj2 lt (pair x y)))))) (omax-x _ _ ) - -ZFP⊆⊗ : {A B : HOD} {x : Ordinal} → odef (ZFP A B) x → odef (A ⊗ B) x -ZFP⊆⊗ {A} {B} {px} ⟪ (pair x y) , p2 ⟫ = product→ (d→∋ A (proj1 (p2 (pair x y) ))) (d→∋ B (proj2 (p2 (pair x y) ))) - --- axiom of choice required --- ⊗⊆ZFP : {A B x : HOD} → ( A ⊗ B ) ∋ x → def ZFProduct (& x) --- ⊗⊆ZFP {A} {B} {x} lt = subst (λ k → ord-pair (& k )) {!!} op-cons -
--- a/OrdUtil.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,287 +0,0 @@ -open import Level -open import Ordinals -module OrdUtil {n : Level} (O : Ordinals {n} ) where - -open import logic -open import nat -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -open import Data.Empty -open import Relation.Binary.PropositionalEquality -open import Relation.Nullary -open import Relation.Binary hiding (_⇔_) - -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext - -o<-dom : { x y : Ordinal } → x o< y → Ordinal -o<-dom {x} _ = x - -o<-cod : { x y : Ordinal } → x o< y → Ordinal -o<-cod {_} {y} _ = y - -o<-subst : {Z X z x : Ordinal } → Z o< X → Z ≡ z → X ≡ x → z o< x -o<-subst df refl refl = df - -o<¬≡ : { ox oy : Ordinal } → ox ≡ oy → ox o< oy → ⊥ -o<¬≡ {ox} {oy} eq lt with trio< ox oy -o<¬≡ {ox} {oy} eq lt | tri< a ¬b ¬c = ¬b eq -o<¬≡ {ox} {oy} eq lt | tri≈ ¬a b ¬c = ¬a lt -o<¬≡ {ox} {oy} eq lt | tri> ¬a ¬b c = ¬b eq - -o<> : {x y : Ordinal } → y o< x → x o< y → ⊥ -o<> {ox} {oy} lt tl with trio< ox oy -o<> {ox} {oy} lt tl | tri< a ¬b ¬c = ¬c lt -o<> {ox} {oy} lt tl | tri≈ ¬a b ¬c = ¬a tl -o<> {ox} {oy} lt tl | tri> ¬a ¬b c = ¬a tl - -osuc-< : { x y : Ordinal } → y o< osuc x → x o< y → ⊥ -osuc-< {x} {y} y<ox x<y with osuc-≡< y<ox -osuc-< {x} {y} y<ox x<y | case1 refl = o<¬≡ refl x<y -osuc-< {x} {y} y<ox x<y | case2 y<x = o<> x<y y<x - -osucc : {ox oy : Ordinal } → oy o< ox → osuc oy o< osuc ox ----- y < osuc y < x < osuc x ----- y < osuc y = x < osuc x ----- y < osuc y > x < osuc x -> y = x ∨ x < y → ⊥ -osucc {ox} {oy} oy<ox with trio< (osuc oy) ox -osucc {ox} {oy} oy<ox | tri< a ¬b ¬c = ordtrans a <-osuc -osucc {ox} {oy} oy<ox | tri≈ ¬a refl ¬c = <-osuc -osucc {ox} {oy} oy<ox | tri> ¬a ¬b c with osuc-≡< c -osucc {ox} {oy} oy<ox | tri> ¬a ¬b c | case1 eq = ⊥-elim (o<¬≡ (sym eq) oy<ox) -osucc {ox} {oy} oy<ox | tri> ¬a ¬b c | case2 lt = ⊥-elim (o<> lt oy<ox) - -osucprev : {ox oy : Ordinal } → osuc oy o< osuc ox → oy o< ox -osucprev {ox} {oy} oy<ox with trio< oy ox -osucprev {ox} {oy} oy<ox | tri< a ¬b ¬c = a -osucprev {ox} {oy} oy<ox | tri≈ ¬a b ¬c = ⊥-elim (o<¬≡ (cong (λ k → osuc k) b) oy<ox ) -osucprev {ox} {oy} oy<ox | tri> ¬a ¬b c = ⊥-elim (o<> (osucc c) oy<ox ) - -open _∧_ - -osuc2 : ( x y : Ordinal ) → ( osuc x o< osuc (osuc y )) ⇔ (x o< osuc y) -proj2 (osuc2 x y) lt = osucc lt -proj1 (osuc2 x y) ox<ooy with osuc-≡< ox<ooy -proj1 (osuc2 x y) ox<ooy | case1 ox=oy = o<-subst <-osuc refl ox=oy -proj1 (osuc2 x y) ox<ooy | case2 ox<oy = ordtrans <-osuc ox<oy - -_o≤_ : Ordinal → Ordinal → Set n -a o≤ b = a o< osuc b -- (a ≡ b) ∨ ( a o< b ) - - -xo<ab : {oa ob : Ordinal } → ( {ox : Ordinal } → ox o< oa → ox o< ob ) → oa o< osuc ob -xo<ab {oa} {ob} a→b with trio< oa ob -xo<ab {oa} {ob} a→b | tri< a ¬b ¬c = ordtrans a <-osuc -xo<ab {oa} {ob} a→b | tri≈ ¬a refl ¬c = <-osuc -xo<ab {oa} {ob} a→b | tri> ¬a ¬b c = ⊥-elim ( o<¬≡ refl (a→b c ) ) - -maxα : Ordinal → Ordinal → Ordinal -maxα x y with trio< x y -maxα x y | tri< a ¬b ¬c = y -maxα x y | tri> ¬a ¬b c = x -maxα x y | tri≈ ¬a refl ¬c = x - -omin : Ordinal → Ordinal → Ordinal -omin x y with trio< x y -omin x y | tri< a ¬b ¬c = x -omin x y | tri> ¬a ¬b c = y -omin x y | tri≈ ¬a refl ¬c = x - -min1 : {x y z : Ordinal } → z o< x → z o< y → z o< omin x y -min1 {x} {y} {z} z<x z<y with trio< x y -min1 {x} {y} {z} z<x z<y | tri< a ¬b ¬c = z<x -min1 {x} {y} {z} z<x z<y | tri≈ ¬a refl ¬c = z<x -min1 {x} {y} {z} z<x z<y | tri> ¬a ¬b c = z<y - --- --- max ( osuc x , osuc y ) --- - -omax : ( x y : Ordinal ) → Ordinal -omax x y with trio< x y -omax x y | tri< a ¬b ¬c = osuc y -omax x y | tri> ¬a ¬b c = osuc x -omax x y | tri≈ ¬a refl ¬c = osuc x - -omax< : ( x y : Ordinal ) → x o< y → osuc y ≡ omax x y -omax< x y lt with trio< x y -omax< x y lt | tri< a ¬b ¬c = refl -omax< x y lt | tri≈ ¬a b ¬c = ⊥-elim (¬a lt ) -omax< x y lt | tri> ¬a ¬b c = ⊥-elim (¬a lt ) - -omax≤ : ( x y : Ordinal ) → x o≤ y → osuc y ≡ omax x y -omax≤ x y le with trio< x y -omax≤ x y le | tri< a ¬b ¬c = refl -omax≤ x y le | tri≈ ¬a refl ¬c = refl -omax≤ x y le | tri> ¬a ¬b c with osuc-≡< le -omax≤ x y le | tri> ¬a ¬b c | case1 eq = ⊥-elim (¬b eq) -omax≤ x y le | tri> ¬a ¬b c | case2 x<y = ⊥-elim (¬a x<y) - -omax≡ : ( x y : Ordinal ) → x ≡ y → osuc y ≡ omax x y -omax≡ x y eq with trio< x y -omax≡ x y eq | tri< a ¬b ¬c = ⊥-elim (¬b eq ) -omax≡ x y eq | tri≈ ¬a refl ¬c = refl -omax≡ x y eq | tri> ¬a ¬b c = ⊥-elim (¬b eq ) - -omax-x : ( x y : Ordinal ) → x o< omax x y -omax-x x y with trio< x y -omax-x x y | tri< a ¬b ¬c = ordtrans a <-osuc -omax-x x y | tri> ¬a ¬b c = <-osuc -omax-x x y | tri≈ ¬a refl ¬c = <-osuc - -omax-y : ( x y : Ordinal ) → y o< omax x y -omax-y x y with trio< x y -omax-y x y | tri< a ¬b ¬c = <-osuc -omax-y x y | tri> ¬a ¬b c = ordtrans c <-osuc -omax-y x y | tri≈ ¬a refl ¬c = <-osuc - -omxx : ( x : Ordinal ) → omax x x ≡ osuc x -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 - -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 )) - -open _∧_ - -o≤-refl : { i j : Ordinal } → i ≡ j → i o≤ j -o≤-refl {i} {j} eq = subst (λ k → i o< osuc k ) eq <-osuc -OrdTrans : Transitive _o≤_ -OrdTrans a≤b b≤c with osuc-≡< a≤b | osuc-≡< b≤c -OrdTrans a≤b b≤c | case1 refl | case1 refl = <-osuc -OrdTrans a≤b b≤c | case1 refl | case2 a≤c = ordtrans a≤c <-osuc -OrdTrans a≤b b≤c | case2 a≤c | case1 refl = ordtrans a≤c <-osuc -OrdTrans a≤b b≤c | case2 a<b | case2 b<c = ordtrans (ordtrans a<b b<c) <-osuc - -OrdPreorder : Preorder n n n -OrdPreorder = record { Carrier = Ordinal - ; _≈_ = _≡_ - ; _∼_ = _o≤_ - ; isPreorder = record { - isEquivalence = record { refl = refl ; sym = sym ; trans = trans } - ; reflexive = o≤-refl - ; trans = OrdTrans - } - } - -FExists : {m l : Level} → ( ψ : Ordinal → Set m ) - → {p : Set l} ( P : { y : Ordinal } → ψ y → ¬ p ) - → (exists : ¬ (∀ y → ¬ ( ψ y ) )) - → ¬ p -FExists {m} {l} ψ {p} P = contra-position ( λ p y ψy → P {y} ψy p ) - -nexto∅ : {x : Ordinal} → o∅ o< next x -nexto∅ {x} with trio< o∅ x -nexto∅ {x} | tri< a ¬b ¬c = ordtrans a x<nx -nexto∅ {x} | tri≈ ¬a b ¬c = subst (λ k → k o< next x) (sym b) x<nx -nexto∅ {x} | tri> ¬a ¬b c = ⊥-elim ( ¬x<0 c ) - -next< : {x y z : Ordinal} → x o< next z → y o< next x → y o< next z -next< {x} {y} {z} x<nz y<nx with trio< y (next z) -next< {x} {y} {z} x<nz y<nx | tri< a ¬b ¬c = a -next< {x} {y} {z} x<nz y<nx | tri≈ ¬a b ¬c = ⊥-elim (¬nx<nx x<nz (subst (λ k → k o< next x) b y<nx) - (λ w nz=ow → o<¬≡ nz=ow (subst₂ (λ j k → j o< k ) (sym nz=ow) nz=ow (osuc<nx (subst (λ k → w o< k ) (sym nz=ow) <-osuc) )))) -next< {x} {y} {z} x<nz y<nx | tri> ¬a ¬b c = ⊥-elim (¬nx<nx x<nz (ordtrans c y<nx ) - (λ w nz=ow → o<¬≡ (sym nz=ow) (osuc<nx (subst (λ k → w o< k ) (sym nz=ow) <-osuc )))) - -osuc< : {x y : Ordinal} → osuc x ≡ y → x o< y -osuc< {x} {y} refl = <-osuc - -nexto=n : {x y : Ordinal} → x o< next (osuc y) → x o< next y -nexto=n {x} {y} x<noy = next< (osuc<nx x<nx) x<noy - -nexto≡ : {x : Ordinal} → next x ≡ next (osuc x) -nexto≡ {x} with trio< (next x) (next (osuc x) ) --- next x o< next (osuc x ) -> osuc x o< next x o< next (osuc x) -> next x ≡ osuc z -> z o o< next x -> osuc z o< next x -> next x o< next x -nexto≡ {x} | tri< a ¬b ¬c = ⊥-elim (¬nx<nx (osuc<nx x<nx ) a - (λ z eq → o<¬≡ (sym eq) (osuc<nx (osuc< (sym eq))))) -nexto≡ {x} | tri≈ ¬a b ¬c = b --- next (osuc x) o< next x -> osuc x o< next (osuc x) o< next x -> next (osuc x) ≡ osuc z -> z o o< next (osuc x) ... -nexto≡ {x} | tri> ¬a ¬b c = ⊥-elim (¬nx<nx (ordtrans <-osuc x<nx) c - (λ z eq → o<¬≡ (sym eq) (osuc<nx (osuc< (sym eq))))) - -next-is-limit : {x y : Ordinal} → ¬ (next x ≡ osuc y) -next-is-limit {x} {y} eq = o<¬≡ (sym eq) (osuc<nx y<nx) where - y<nx : y o< next x - y<nx = osuc< (sym eq) - -omax<next : {x y : Ordinal} → x o< y → omax x y o< next y -omax<next {x} {y} x<y = subst (λ k → k o< next y ) (omax< _ _ x<y ) (osuc<nx x<nx) - -x<ny→≡next : {x y : Ordinal} → x o< y → y o< next x → next x ≡ next y -x<ny→≡next {x} {y} x<y y<nx with trio< (next x) (next y) -x<ny→≡next {x} {y} x<y y<nx | tri< a ¬b ¬c = -- x < y < next x < next y ∧ next x = osuc z - ⊥-elim ( ¬nx<nx y<nx a (λ z eq → o<¬≡ (sym eq) (osuc<nx (subst (λ k → z o< k ) (sym eq) <-osuc )))) -x<ny→≡next {x} {y} x<y y<nx | tri≈ ¬a b ¬c = b -x<ny→≡next {x} {y} x<y y<nx | tri> ¬a ¬b c = -- x < y < next y < next x - ⊥-elim ( ¬nx<nx (ordtrans x<y x<nx) c (λ z eq → o<¬≡ (sym eq) (osuc<nx (subst (λ k → z o< k ) (sym eq) <-osuc )))) - -≤next : {x y : Ordinal} → x o≤ y → next x o≤ next y -≤next {x} {y} x≤y with trio< (next x) y -≤next {x} {y} x≤y | tri< a ¬b ¬c = ordtrans a (ordtrans x<nx <-osuc ) -≤next {x} {y} x≤y | tri≈ ¬a refl ¬c = (ordtrans x<nx <-osuc ) -≤next {x} {y} x≤y | tri> ¬a ¬b c with osuc-≡< x≤y -≤next {x} {y} x≤y | tri> ¬a ¬b c | case1 refl = o≤-refl refl -- x = y < next x -≤next {x} {y} x≤y | tri> ¬a ¬b c | case2 x<y = o≤-refl (x<ny→≡next x<y c) -- x ≤ y < next x - -x<ny→≤next : {x y : Ordinal} → x o< next y → next x o≤ next y -x<ny→≤next {x} {y} x<ny with trio< x y -x<ny→≤next {x} {y} x<ny | tri< a ¬b ¬c = ≤next (ordtrans a <-osuc ) -x<ny→≤next {x} {y} x<ny | tri≈ ¬a refl ¬c = o≤-refl refl -x<ny→≤next {x} {y} x<ny | tri> ¬a ¬b c = o≤-refl (sym ( x<ny→≡next c x<ny )) - -omax<nomax : {x y : Ordinal} → omax x y o< next (omax x y ) -omax<nomax {x} {y} with trio< x y -omax<nomax {x} {y} | tri< a ¬b ¬c = subst (λ k → osuc y o< k ) nexto≡ (osuc<nx x<nx ) -omax<nomax {x} {y} | tri≈ ¬a refl ¬c = subst (λ k → osuc x o< k ) nexto≡ (osuc<nx x<nx ) -omax<nomax {x} {y} | tri> ¬a ¬b c = subst (λ k → osuc x o< k ) nexto≡ (osuc<nx x<nx ) - -omax<nx : {x y z : Ordinal} → x o< next z → y o< next z → omax x y o< next z -omax<nx {x} {y} {z} x<nz y<nz with trio< x y -omax<nx {x} {y} {z} x<nz y<nz | tri< a ¬b ¬c = osuc<nx y<nz -omax<nx {x} {y} {z} x<nz y<nz | tri≈ ¬a refl ¬c = osuc<nx y<nz -omax<nx {x} {y} {z} x<nz y<nz | tri> ¬a ¬b c = osuc<nx x<nz - -nn<omax : {x nx ny : Ordinal} → x o< next nx → x o< next ny → x o< next (omax nx ny) -nn<omax {x} {nx} {ny} xnx xny with trio< nx ny -nn<omax {x} {nx} {ny} xnx xny | tri< a ¬b ¬c = subst (λ k → x o< k ) nexto≡ xny -nn<omax {x} {nx} {ny} xnx xny | tri≈ ¬a refl ¬c = subst (λ k → x o< k ) nexto≡ xny -nn<omax {x} {nx} {ny} xnx xny | tri> ¬a ¬b c = subst (λ k → x o< k ) nexto≡ xnx - -record OrdinalSubset (maxordinal : Ordinal) : Set (suc n) where - field - os→ : (x : Ordinal) → x o< maxordinal → Ordinal - os← : Ordinal → Ordinal - os←limit : (x : Ordinal) → os← x o< maxordinal - os-iso← : {x : Ordinal} → os→ (os← x) (os←limit x) ≡ x - os-iso→ : {x : Ordinal} → (lt : x o< maxordinal ) → os← (os→ x lt) ≡ x - -module o≤-Reasoning {n : Level} (O : Ordinals {n} ) where - - -- open inOrdinal O - - resp-o< : _o<_ Respects₂ _≡_ - resp-o< = resp₂ _o<_ - - trans1 : {i j k : Ordinal} → i o< j → j o< osuc k → i o< k - trans1 {i} {j} {k} i<j j<ok with osuc-≡< j<ok - trans1 {i} {j} {k} i<j j<ok | case1 refl = i<j - trans1 {i} {j} {k} i<j j<ok | case2 j<k = ordtrans i<j j<k - - trans2 : {i j k : Ordinal} → i o< osuc j → j o< k → i o< k - trans2 {i} {j} {k} i<oj j<k with osuc-≡< i<oj - trans2 {i} {j} {k} i<oj j<k | case1 refl = j<k - trans2 {i} {j} {k} i<oj j<k | case2 i<j = ordtrans i<j j<k - - open import Relation.Binary.Reasoning.Base.Triple - (Preorder.isPreorder OrdPreorder) - ordtrans --<-trans - (resp₂ _o<_) --(resp₂ _<_) - (λ x → ordtrans x <-osuc ) --<⇒≤ - trans1 --<-transˡ - trans2 --<-transʳ - public - -- hiding (_≈⟨_⟩_) -
--- a/Ordinals.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,62 +0,0 @@ -open import Level -module Ordinals where - -open import zf - -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -open import Data.Empty -open import Relation.Binary.PropositionalEquality -open import logic -open import nat -open import Data.Unit using ( ⊤ ) -open import Relation.Nullary -open import Relation.Binary -open import Relation.Binary.Core - -record Oprev {n : Level} (ord : Set n) (osuc : ord → ord ) (x : ord ) : Set (suc n) where - field - oprev : ord - oprev=x : osuc oprev ≡ x - -record IsOrdinals {n : Level} (ord : Set n) (o∅ : ord ) (osuc : ord → ord ) (_o<_ : ord → ord → Set n) (next : ord → ord ) : Set (suc (suc n)) where - field - ordtrans : {x y z : ord } → x o< y → y o< z → x o< z - trio< : Trichotomous {n} _≡_ _o<_ - ¬x<0 : { x : ord } → ¬ ( x o< o∅ ) - <-osuc : { x : ord } → x o< osuc x - osuc-≡< : { a x : ord } → x o< osuc a → (x ≡ a ) ∨ (x o< a) - Oprev-p : ( x : ord ) → Dec ( Oprev ord osuc x ) - TransFinite : { ψ : ord → Set (suc n) } - → ( (x : ord) → ( (y : ord ) → y o< x → ψ y ) → ψ x ) - → ∀ (x : ord) → ψ x - - -record IsNext {n : Level } (ord : Set n) (o∅ : ord ) (osuc : ord → ord ) (_o<_ : ord → ord → Set n) (next : ord → ord ) : Set (suc (suc n)) where - field - x<nx : { y : ord } → (y o< next y ) - osuc<nx : { x y : ord } → x o< next y → osuc x o< next y - ¬nx<nx : {x y : ord} → y o< x → x o< next y → ¬ ((z : ord) → ¬ (x ≡ osuc z)) - -record Ordinals {n : Level} : Set (suc (suc n)) where - field - Ordinal : Set n - o∅ : Ordinal - osuc : Ordinal → Ordinal - _o<_ : Ordinal → Ordinal → Set n - next : Ordinal → Ordinal - isOrdinal : IsOrdinals Ordinal o∅ osuc _o<_ next - isNext : IsNext Ordinal o∅ osuc _o<_ next - -module inOrdinal {n : Level} (O : Ordinals {n} ) where - - open Ordinals O - open IsOrdinals isOrdinal - open IsNext isNext - - TransFinite0 : { ψ : Ordinal → Set n } - → ( (x : Ordinal) → ( (y : Ordinal ) → y o< x → ψ y ) → ψ x ) - → ∀ (x : Ordinal) → ψ x - TransFinite0 {ψ} ind x = lower (TransFinite {λ y → Lift (suc n) ( ψ y)} ind1 x) where - ind1 : (z : Ordinal) → ((y : Ordinal) → y o< z → Lift (suc n) (ψ y)) → Lift (suc n) (ψ z) - ind1 z prev = lift (ind z (λ y y<z → lower (prev y y<z ) )) -
--- a/README.md Sun Dec 20 12:37:07 2020 +0900 +++ b/README.md Mon Dec 21 10:23:37 2020 +0900 @@ -78,9 +78,9 @@ _∋_ A x = def (od A) ( od→ord x ) ``` -In ψ : Ordinal → Set, if A is a record { def = λ x → ψ x } , then +In ψ : Ordinal → Set, if A is a record { od = { def = λ x → ψ x } ...} , then - A x = def A ( od→ord x ) = ψ (od→ord x) + A ∋ x = def (od A) ( od→ord x ) = ψ (od→ord x) They say the existing of the mappings can be proved in Classical Set Theory, but we simply assumes these non constructively.
--- a/Topology.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,109 +0,0 @@ -open import Level -open import Ordinals -module Topology {n : Level } (O : Ordinals {n}) where - -open import zf -open import logic -open _∧_ -open _∨_ -open Bool - -import OD -open import Relation.Nullary -open import Data.Empty -open import Relation.Binary.Core -open import Relation.Binary.PropositionalEquality -import BAlgbra -open BAlgbra O -open inOrdinal O -open OD O -open OD.OD -open ODAxiom odAxiom -import OrdUtil -import ODUtil -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -open OrdUtil O -open ODUtil O - -import ODC -open ODC O - -open import filter - -record Toplogy ( L : HOD ) : Set (suc n) where - field - OS : HOD - OS⊆PL : OS ⊆ Power L - o∪ : { P : HOD } → P ⊆ OS → OS ∋ Union P - o∩ : { p q : HOD } → OS ∋ p → OS ∋ q → OS ∋ (p ∩ q) - -open Toplogy - -record _covers_ ( P q : HOD ) : Set (suc n) where - field - cover : {x : HOD} → q ∋ x → HOD - P∋cover : {x : HOD} → {lt : q ∋ x} → P ∋ cover lt - isCover : {x : HOD} → {lt : q ∋ x} → cover lt ∋ x - --- Base --- The elements of B cover X ; For any U , V ∈ B and any point x ∈ U ∩ V there is a W ∈ B such that --- W ⊆ U ∩ V and x ∈ W . - -data genTop (P : HOD) : HOD → Set (suc n) where - gi : {x : HOD} → P ∋ x → genTop P x - g∩ : {x y : HOD} → genTop P x → genTop P y → genTop P (x ∩ y) - g∪ : {Q x : HOD} → Q ⊆ P → genTop P (Union Q) - --- Limit point - -record LP ( L S x : HOD ) (top : Toplogy L) (S⊆PL : S ⊆ Power L ) ( S∋x : S ∋ x ) : Set (suc n) where - field - neip : {y : HOD} → OS top ∋ y → y ∋ x → HOD - isNeip : {y : HOD} → (o∋y : OS top ∋ y ) → (y∋x : y ∋ x ) → ¬ ( x ≡ neip o∋y y∋x) ∧ ( y ∋ neip o∋y y∋x ) - --- Finite Intersection Property - -data Finite-∩ (S : HOD) : HOD → Set (suc n) where - fin-∩e : {x : HOD} → S ∋ x → Finite-∩ S x - fin-∩ : {x y : HOD} → Finite-∩ S x → Finite-∩ S y → Finite-∩ S (x ∩ y) - -record FIP ( L P : HOD ) : Set (suc n) where - field - fipS⊆PL : P ⊆ Power L - fip≠φ : { x : HOD } → Finite-∩ P x → ¬ ( x ≡ od∅ ) - --- Compact - -data Finite-∪ (S : HOD) : HOD → Set (suc n) where - fin-∪e : {x : HOD} → S ∋ x → Finite-∪ S x - fin-∪ : {x y : HOD} → Finite-∪ S x → Finite-∪ S y → Finite-∪ S (x ∪ y) - -record Compact ( L P : HOD ) : Set (suc n) where - field - finCover : {X y : HOD} → X covers P → P ∋ y → HOD - isFinCover : {X y : HOD} → (xp : X covers P ) → (P∋y : P ∋ y ) → finCover xp P∋y ∋ y - isFininiteCover : {X y : HOD} → (xp : X covers P ) → (P∋y : P ∋ y ) → Finite-∪ X (finCover xp P∋y ) - --- FIP is Compact - -FIP→Compact : {L P : HOD} → Tolopogy L → FIP L P → Compact L P -FIP→Compact = ? - -Compact→FIP : {L P : HOD} → Tolopogy L → Compact L P → FIP L P -Compact→FIP = ? - --- Product Topology - -_Top⊗_ : {P Q : HOD} → Topology P → Tolopogy Q → Topology ( P ⊗ Q ) -_Top⊗_ = ? - --- existence of Ultra Filter - --- Ultra Filter has limit point - --- FIP is UFL - --- Product of UFL has limit point (Tychonoff) -
--- a/VL.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,86 +0,0 @@ -open import Level -open import Ordinals -module VL {n : Level } (O : Ordinals {n}) where - -open import zf -open import logic -import OD -open import Relation.Nullary -open import Relation.Binary -open import Data.Empty -open import Relation.Binary -open import Relation.Binary.Core -open import Relation.Binary.PropositionalEquality -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -import BAlgbra -open BAlgbra O -open inOrdinal O -import OrdUtil -import ODUtil -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -open OrdUtil O -open ODUtil O - -open OD O -open OD.OD -open ODAxiom odAxiom --- import ODC -open _∧_ -open _∨_ -open Bool -open HOD - --- The cumulative hierarchy --- V 0 := ∅ --- V α + 1 := P ( V α ) --- V α := ⋃ { V β | β < α } - -V : ( β : Ordinal ) → HOD -V β = TransFinite V1 β where - V1 : (x : Ordinal ) → ( ( y : Ordinal) → y o< x → HOD ) → HOD - V1 x V0 with trio< x o∅ - V1 x V0 | tri< a ¬b ¬c = ⊥-elim ( ¬x<0 a) - V1 x V0 | tri≈ ¬a refl ¬c = Ord o∅ - V1 x V0 | tri> ¬a ¬b c with Oprev-p x - V1 x V0 | tri> ¬a ¬b c | yes p = Power ( V0 (Oprev.oprev p ) (subst (λ k → _ o< k) (Oprev.oprev=x p) <-osuc )) - V1 x V0 | tri> ¬a ¬b c | no ¬p = - record { od = record { def = λ y → (y o< x ) ∧ ((lt : y o< x ) → odef (V0 y lt) x ) } ; odmax = x; <odmax = λ {x} lt → proj1 lt } - --- --- L ⊆ HOD ⊆ V --- --- HOD=V : (x : HOD) → V (odmax x) ∋ x --- HOD=V x = {!!} - --- --- Definition for Power Set --- -record Definitions : Set (suc n) where - field - Definition : HOD → HOD - -open Definitions - -Df : Definitions → (x : HOD) → HOD -Df D x = Power x ∩ Definition D x - --- The constructible Sets --- L 0 := ∅ --- L α + 1 := Df (L α) -- Definable Power Set --- V α := ⋃ { L β | β < α } - -L : ( β : Ordinal ) → Definitions → HOD -L β D = TransFinite L1 β where - L1 : (x : Ordinal ) → ( ( y : Ordinal) → y o< x → HOD ) → HOD - L1 x L0 with trio< x o∅ - L1 x L0 | tri< a ¬b ¬c = ⊥-elim ( ¬x<0 a) - L1 x L0 | tri≈ ¬a refl ¬c = Ord o∅ - L1 x L0 | tri> ¬a ¬b c with Oprev-p x - L1 x L0 | tri> ¬a ¬b c | yes p = Df D ( L0 (Oprev.oprev p ) (subst (λ k → _ o< k) (Oprev.oprev=x p) <-osuc )) - L1 x L0 | tri> ¬a ¬b c | no ¬p = - record { od = record { def = λ y → (y o< x ) ∧ ((lt : y o< x ) → odef (L0 y lt) x ) } ; odmax = x; <odmax = λ {x} lt → proj1 lt } - -V=L0 : Set (suc n) -V=L0 = (x : Ordinal) → V x ≡ L x record { Definition = λ y → y }
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ZF.agda-lib Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,3 @@ +name: ZF +depend: standard-library +include: src
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ZF.agda-pkg Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,13 @@ +# File generated by Agda-Pkg +name: ZF +version: v0.0.1 + + +homepage: https://ie.u-ryukyu.ac.jp/~kono +license: MIT +license-file: LICENSE.md +source-repository: https://github.com/shinji-kono/zf-in-agda +tested-with: 2.6.2-102d9c8 +description: ZF is an Agda library ... + +# End
--- a/cardinal.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,95 +0,0 @@ -open import Level -open import Ordinals -module cardinal {n : Level } (O : Ordinals {n}) where - -open import zf -open import logic -import OD -import ODC -import OPair -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -open import Relation.Binary.PropositionalEquality -open import Data.Nat.Properties -open import Data.Empty -open import Relation.Nullary -open import Relation.Binary -open import Relation.Binary.Core - -open inOrdinal O -open OD O -open OD.OD -open OPair O -open ODAxiom odAxiom - -import OrdUtil -import ODUtil -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -open OrdUtil O -open ODUtil O - - -open _∧_ -open _∨_ -open Bool -open _==_ - -open HOD - --- _⊗_ : (A B : HOD) → HOD --- A ⊗ B = Union ( Replace B (λ b → Replace A (λ a → < a , b > ) )) - -Func : OD -Func = record { def = λ x → def ZFProduct x - ∧ ( (a b c : Ordinal) → odef (* x) (& < * a , * b >) ∧ odef (* x) (& < * a , * c >) → b ≡ c ) } - -FuncP : ( A B : HOD ) → HOD -FuncP A B = record { od = record { def = λ x → odef (ZFP A B) x - ∧ ( (x : Ordinal ) (p q : odef (ZFP A B ) x ) → pi1 (proj1 p) ≡ pi1 (proj1 q) → pi2 (proj1 p) ≡ pi2 (proj1 q) ) } - ; odmax = odmax (ZFP A B) ; <odmax = λ lt → <odmax (ZFP A B) (proj1 lt) } - -record Injection (A B : Ordinal ) : Set n where - field - i→ : (x : Ordinal ) → odef (* A) x → Ordinal - iB : (x : Ordinal ) → ( lt : odef (* A) x ) → odef (* B) ( i→ x lt ) - iiso : (x y : Ordinal ) → ( ltx : odef (* A) x ) ( lty : odef (* A) y ) → i→ x ltx ≡ i→ y lty → x ≡ y - -record Bijection (A B : Ordinal ) : Set n where - field - fun← : (x : Ordinal ) → odef (* A) x → Ordinal - fun→ : (x : Ordinal ) → odef (* B) x → Ordinal - funB : (x : Ordinal ) → ( lt : odef (* A) x ) → odef (* B) ( fun← x lt ) - funA : (x : Ordinal ) → ( lt : odef (* B) x ) → odef (* A) ( fun→ x lt ) - fiso← : (x : Ordinal ) → ( lt : odef (* B) x ) → fun← ( fun→ x lt ) ( funA x lt ) ≡ x - fiso→ : (x : Ordinal ) → ( lt : odef (* A) x ) → fun→ ( fun← x lt ) ( funB x lt ) ≡ x - -Bernstein : {a b : Ordinal } → Injection a b → Injection b a → Bijection a b -Bernstein = {!!} - -_=c=_ : ( A B : HOD ) → Set n -A =c= B = Bijection ( & A ) ( & B ) - -_c<_ : ( A B : HOD ) → Set n -A c< B = ¬ ( Injection (& A) (& B) ) - -Card : OD -Card = record { def = λ x → (a : Ordinal) → a o< x → ¬ Bijection a x } - -record Cardinal (a : Ordinal ) : Set (suc n) where - field - card : Ordinal - ciso : Bijection a card - cmax : (x : Ordinal) → card o< x → ¬ Bijection a x - -Cardinal∈ : { s : HOD } → { t : Ordinal } → Ord t ∋ s → s c< Ord t -Cardinal∈ = {!!} - -Cardinal⊆ : { s t : HOD } → s ⊆ t → ( s c< t ) ∨ ( s =c= t ) -Cardinal⊆ = {!!} - -Cantor1 : { u : HOD } → u c< Power u -Cantor1 = {!!} - -Cantor2 : { u : HOD } → ¬ ( u =c= Power u ) -Cantor2 = {!!}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/fig/Sets.txt Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,3 @@ +Ordinal V L OD Set + + \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/fig/zf-record.html Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,55 @@ +<html> +<META HTTP-EQUIV="Content-Type" CONTENT="text/html; charset=UTF-8"> + +<pre> + record ZF {n m : Level } : Set (suc (n ⊔ m)) where + field + ZFSet : Set n + _∋_ : ( A x : ZFSet ) → Set m + _≈_ : ( A B : ZFSet ) → Set m + ∅ : ZFSet + _,_ : ( A B : ZFSet ) → ZFSet + Union : ( A : ZFSet ) → ZFSet + Power : ( A : ZFSet ) → ZFSet + Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet + Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet + infinite : ZFSet + isZF : IsZF ZFSet _∋_ _≈_ ∅ _,_ Union Power Select Replace infinite + record IsZF {n m : Level } + (ZFSet : Set n) + (_∋_ : ( A x : ZFSet ) → Set m) + (_≈_ : Rel ZFSet m) + (∅ : ZFSet) + (_,_ : ( A B : ZFSet ) → ZFSet) + (Union : ( A : ZFSet ) → ZFSet) + (Power : ( A : ZFSet ) → ZFSet) + (Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet ) + (Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet ) + (infinite : ZFSet) + : Set (suc (n ⊔ m)) where + field + isEquivalence : IsEquivalence {n} {m} {ZFSet} _≈_ + pair→ : ( x y t : ZFSet ) → (x , y) ∋ t → ( t ≈ x ) ∨ ( t ≈ y ) + pair← : ( x y t : ZFSet ) → ( t ≈ x ) ∨ ( t ≈ y ) → (x , y) ∋ t + union→ : ( X z u : ZFSet ) → ( X ∋ u ) ∧ (u ∋ z ) → Union X ∋ z + union← : ( X z : ZFSet ) → (X∋z : Union X ∋ z ) → ¬ ( (u : ZFSet ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) + empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x ) + power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → t ∋ x → ¬ ¬ ( A ∋ x ) -- _⊆_ t A {x} + power← : ∀( A t : ZFSet ) → ( ∀ {x} → _⊆_ t A {x}) → Power A ∋ t + extensionality : { A B w : ZFSet } → ( (z : ZFSet) → ( A ∋ z ) ⇔ (B ∋ z) ) → ( A ∈ w ⇔ B ∈ w ) + regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) ) + minimal : (x : ZFSet ) → ¬ (x ≈ ∅) → ZFSet + regularity : ∀( x : ZFSet ) → (not : ¬ (x ≈ ∅)) → ( minimal x not ∈ x ∧ ( minimal x not ∩ x ≈ ∅ ) ) + ε-induction : { ψ : ZFSet → Set m} + → ( {x : ZFSet } → ({ y : ZFSet } → x ∋ y → ψ y ) → ψ x ) + → (x : ZFSet ) → ψ x + infinity∅ : ∅ ∈ infinite + infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ { x }) ∈ infinite + selection : { ψ : ZFSet → Set m } → ∀ { X y : ZFSet } → ( ( y ∈ X ) ∧ ψ y ) ⇔ (y ∈ Select X ψ ) + replacement← : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → x ∈ X → ψ x ∈ Replace X ψ + replacement→ : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → ( lt : x ∈ Replace X ψ ) → ¬ ( ∀ (y : ZFSet) → ¬ ( x ≈ ψ y ) ) + choice-func : (X : ZFSet ) → {x : ZFSet } → ¬ ( x ≈ ∅ ) → ( X ∋ x ) → ZFSet + choice : (X : ZFSet ) → {A : ZFSet } → ( X∋A : X ∋ A ) → (not : ¬ ( A ≈ ∅ )) → A ∋ choice-func X not X∋A +</div> +<hr/> Shinji KONO <kono@ie.u-ryukyu.ac.jp> / Fri Jan 10 12:24:29 2020 +</body></html>
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/fig/zf-record.ind Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,50 @@ + record ZF {n m : Level } : Set (suc (n ⊔ m)) where + field + ZFSet : Set n + _∋_ : ( A x : ZFSet ) → Set m + _≈_ : ( A B : ZFSet ) → Set m + ∅ : ZFSet + _,_ : ( A B : ZFSet ) → ZFSet + Union : ( A : ZFSet ) → ZFSet + Power : ( A : ZFSet ) → ZFSet + Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet + Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet + infinite : ZFSet + isZF : IsZF ZFSet _∋_ _≈_ ∅ _,_ Union Power Select Replace infinite + + record IsZF {n m : Level } + (ZFSet : Set n) + (_∋_ : ( A x : ZFSet ) → Set m) + (_≈_ : Rel ZFSet m) + (∅ : ZFSet) + (_,_ : ( A B : ZFSet ) → ZFSet) + (Union : ( A : ZFSet ) → ZFSet) + (Power : ( A : ZFSet ) → ZFSet) + (Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet ) + (Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet ) + (infinite : ZFSet) + : Set (suc (n ⊔ m)) where + field + isEquivalence : IsEquivalence {n} {m} {ZFSet} _≈_ + pair→ : ( x y t : ZFSet ) → (x , y) ∋ t → ( t ≈ x ) ∨ ( t ≈ y ) + pair← : ( x y t : ZFSet ) → ( t ≈ x ) ∨ ( t ≈ y ) → (x , y) ∋ t + union→ : ( X z u : ZFSet ) → ( X ∋ u ) ∧ (u ∋ z ) → Union X ∋ z + union← : ( X z : ZFSet ) → (X∋z : Union X ∋ z ) → ¬ ( (u : ZFSet ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) + empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x ) + power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → t ∋ x → ¬ ¬ ( A ∋ x ) -- _⊆_ t A {x} + power← : ∀( A t : ZFSet ) → ( ∀ {x} → _⊆_ t A {x}) → Power A ∋ t + extensionality : { A B w : ZFSet } → ( (z : ZFSet) → ( A ∋ z ) ⇔ (B ∋ z) ) → ( A ∈ w ⇔ B ∈ w ) + regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) ) + minimal : (x : ZFSet ) → ¬ (x ≈ ∅) → ZFSet + regularity : ∀( x : ZFSet ) → (not : ¬ (x ≈ ∅)) → ( minimal x not ∈ x ∧ ( minimal x not ∩ x ≈ ∅ ) ) + ε-induction : { ψ : ZFSet → Set m} + → ( {x : ZFSet } → ({ y : ZFSet } → x ∋ y → ψ y ) → ψ x ) + → (x : ZFSet ) → ψ x + infinity∅ : ∅ ∈ infinite + infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ { x }) ∈ infinite + selection : { ψ : ZFSet → Set m } → ∀ { X y : ZFSet } → ( ( y ∈ X ) ∧ ψ y ) ⇔ (y ∈ Select X ψ ) + replacement← : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → x ∈ X → ψ x ∈ Replace X ψ + replacement→ : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → ( lt : x ∈ Replace X ψ ) → ¬ ( ∀ (y : ZFSet) → ¬ ( x ≈ ψ y ) ) + choice-func : (X : ZFSet ) → {x : ZFSet } → ¬ ( x ≈ ∅ ) → ( X ∋ x ) → ZFSet + choice : (X : ZFSet ) → {A : ZFSet } → ( X∋A : X ∋ A ) → (not : ¬ ( A ≈ ∅ )) → A ∋ choice-func X not X∋A +
--- a/filter.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,195 +0,0 @@ -open import Level -open import Ordinals -module filter {n : Level } (O : Ordinals {n}) where - -open import zf -open import logic -import OD - -open import Relation.Nullary -open import Data.Empty -open import Relation.Binary.Core -open import Relation.Binary.PropositionalEquality -import BAlgbra - -open BAlgbra O - -open inOrdinal O -open OD O -open OD.OD -open ODAxiom odAxiom - -import OrdUtil -import ODUtil -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -open OrdUtil O -open ODUtil O - - -import ODC -open ODC O - -open _∧_ -open _∨_ -open Bool - --- Kunen p.76 and p.53, we use ⊆ -record Filter ( L : HOD ) : Set (suc n) where - field - filter : HOD - f⊆PL : filter ⊆ Power L - filter1 : { p q : HOD } → q ⊆ L → filter ∋ p → p ⊆ q → filter ∋ q - filter2 : { p q : HOD } → filter ∋ p → filter ∋ q → filter ∋ (p ∩ q) - -open Filter - -record prime-filter { L : HOD } (P : Filter L) : Set (suc (suc n)) where - field - proper : ¬ (filter P ∋ od∅) - prime : {p q : HOD } → filter P ∋ (p ∪ q) → ( filter P ∋ p ) ∨ ( filter P ∋ q ) - -record ultra-filter { L : HOD } (P : Filter L) : Set (suc (suc n)) where - field - proper : ¬ (filter P ∋ od∅) - ultra : {p : HOD } → p ⊆ L → ( filter P ∋ p ) ∨ ( filter P ∋ ( L \ p) ) - -open _⊆_ - -∈-filter : {L p : HOD} → (P : Filter L ) → filter P ∋ p → p ⊆ L -∈-filter {L} {p} P lt = power→⊆ L p ( incl (f⊆PL P) lt ) - -∪-lemma1 : {L p q : HOD } → (p ∪ q) ⊆ L → p ⊆ L -∪-lemma1 {L} {p} {q} lt = record { incl = λ {x} p∋x → incl lt (case1 p∋x) } - -∪-lemma2 : {L p q : HOD } → (p ∪ q) ⊆ L → q ⊆ L -∪-lemma2 {L} {p} {q} lt = record { incl = λ {x} p∋x → incl lt (case2 p∋x) } - -q∩q⊆q : {p q : HOD } → (q ∩ p) ⊆ q -q∩q⊆q = record { incl = λ lt → proj1 lt } - -open HOD - ------ --- --- ultra filter is prime --- - -filter-lemma1 : {L : HOD} → (P : Filter L) → ∀ {p q : HOD } → ultra-filter P → prime-filter P -filter-lemma1 {L} P u = record { - proper = ultra-filter.proper u - ; prime = lemma3 - } where - lemma3 : {p q : HOD} → filter P ∋ (p ∪ q) → ( filter P ∋ p ) ∨ ( filter P ∋ q ) - lemma3 {p} {q} lt with ultra-filter.ultra u (∪-lemma1 (∈-filter P lt) ) - ... | case1 p∈P = case1 p∈P - ... | case2 ¬p∈P = case2 (filter1 P {q ∩ (L \ p)} (∪-lemma2 (∈-filter P lt)) lemma7 lemma8) where - lemma5 : ((p ∪ q ) ∩ (L \ p)) =h= (q ∩ (L \ p)) - lemma5 = record { eq→ = λ {x} lt → ⟪ lemma4 x lt , proj2 lt ⟫ - ; eq← = λ {x} lt → ⟪ case2 (proj1 lt) , proj2 lt ⟫ - } where - lemma4 : (x : Ordinal ) → odef ((p ∪ q) ∩ (L \ p)) x → odef q x - lemma4 x lt with proj1 lt - lemma4 x lt | case1 px = ⊥-elim ( proj2 (proj2 lt) px ) - lemma4 x lt | case2 qx = qx - lemma6 : filter P ∋ ((p ∪ q ) ∩ (L \ p)) - lemma6 = filter2 P lt ¬p∈P - lemma7 : filter P ∋ (q ∩ (L \ p)) - lemma7 = subst (λ k → filter P ∋ k ) (==→o≡ lemma5 ) lemma6 - lemma8 : (q ∩ (L \ p)) ⊆ q - lemma8 = q∩q⊆q - ------ --- --- if Filter contains L, prime filter is ultra --- - -filter-lemma2 : {L : HOD} → (P : Filter L) → filter P ∋ L → prime-filter P → ultra-filter P -filter-lemma2 {L} P f∋L prime = record { - proper = prime-filter.proper prime - ; ultra = λ {p} p⊆L → prime-filter.prime prime (lemma p p⊆L) - } where - open _==_ - p+1-p=1 : {p : HOD} → p ⊆ L → L =h= (p ∪ (L \ p)) - eq→ (p+1-p=1 {p} p⊆L) {x} lt with ODC.decp O (odef p x) - eq→ (p+1-p=1 {p} p⊆L) {x} lt | yes p∋x = case1 p∋x - eq→ (p+1-p=1 {p} p⊆L) {x} lt | no ¬p = case2 ⟪ lt , ¬p ⟫ - eq← (p+1-p=1 {p} p⊆L) {x} ( case1 p∋x ) = subst (λ k → odef L k ) &iso (incl p⊆L ( subst (λ k → odef p k) (sym &iso) p∋x )) - eq← (p+1-p=1 {p} p⊆L) {x} ( case2 ¬p ) = proj1 ¬p - lemma : (p : HOD) → p ⊆ L → filter P ∋ (p ∪ (L \ p)) - lemma p p⊆L = subst (λ k → filter P ∋ k ) (==→o≡ (p+1-p=1 p⊆L)) f∋L - -record Dense (P : HOD ) : Set (suc n) where - field - dense : HOD - d⊆P : dense ⊆ Power P - dense-f : HOD → HOD - dense-d : { p : HOD} → p ⊆ P → dense ∋ dense-f p - dense-p : { p : HOD} → p ⊆ P → p ⊆ (dense-f p) - -record Ideal ( L : HOD ) : Set (suc n) where - field - ideal : HOD - i⊆PL : ideal ⊆ Power L - ideal1 : { p q : HOD } → q ⊆ L → ideal ∋ p → q ⊆ p → ideal ∋ q - ideal2 : { p q : HOD } → ideal ∋ p → ideal ∋ q → ideal ∋ (p ∪ q) - -open Ideal - -proper-ideal : {L : HOD} → (P : Ideal L ) → {p : HOD} → Set n -proper-ideal {L} P {p} = ideal P ∋ od∅ - -prime-ideal : {L : HOD} → Ideal L → ∀ {p q : HOD } → Set n -prime-ideal {L} P {p} {q} = ideal P ∋ ( p ∩ q) → ( ideal P ∋ p ) ∨ ( ideal P ∋ q ) - ----- --- --- Filter/Ideal without ZF --- L have to be a Latice --- - -record F-Filter {n : Level} (L : Set n) (PL : (L → Set n) → Set n) ( _⊆_ : L → L → Set n) (_∩_ : L → L → L ) : Set (suc n) where - field - filter : L → Set n - f⊆P : PL filter - filter1 : { p q : L } → PL (λ x → q ⊆ x ) → filter p → p ⊆ q → filter q - filter2 : { p q : L } → filter p → filter q → filter (p ∩ q) - -Filter-is-F : {L : HOD} → (f : Filter L ) → F-Filter HOD (λ p → (x : HOD) → p x → x ⊆ L ) _⊆_ _∩_ -Filter-is-F {L} f = record { - filter = λ x → Lift (suc n) ((filter f) ∋ x) - ; f⊆P = λ x f∋x → power→⊆ _ _ (incl ( f⊆PL f ) (lower f∋x )) - ; filter1 = λ {p} {q} q⊆L f∋p p⊆q → lift ( filter1 f (q⊆L q refl-⊆) (lower f∋p) p⊆q) - ; filter2 = λ {p} {q} f∋p f∋q → lift ( filter2 f (lower f∋p) (lower f∋q)) - } - -record F-Dense {n : Level} (L : Set n) (PL : (L → Set n) → Set n) ( _⊆_ : L → L → Set n) (_∩_ : L → L → L ) : Set (suc n) where - field - dense : L → Set n - d⊆P : PL dense - dense-f : L → L - dense-d : { p : L} → PL (λ x → p ⊆ x ) → dense ( dense-f p ) - dense-p : { p : L} → PL (λ x → p ⊆ x ) → p ⊆ (dense-f p) - -Dense-is-F : {L : HOD} → (f : Dense L ) → F-Dense HOD (λ p → (x : HOD) → p x → x ⊆ L ) _⊆_ _∩_ -Dense-is-F {L} f = record { - dense = λ x → Lift (suc n) ((dense f) ∋ x) - ; d⊆P = λ x f∋x → power→⊆ _ _ (incl ( d⊆P f ) (lower f∋x )) - ; dense-f = λ x → dense-f f x - ; dense-d = λ {p} d → lift ( dense-d f (d p refl-⊆ ) ) - ; dense-p = λ {p} d → dense-p f (d p refl-⊆) - } where open Dense - - -record GenericFilter (P : HOD) : Set (suc n) where - field - genf : Filter P - generic : (D : Dense P ) → ¬ ( (Dense.dense D ∩ Filter.filter genf ) ≡ od∅ ) - -record F-GenericFilter {n : Level} (L : Set n) (PL : (L → Set n) → Set n) ( _⊆_ : L → L → Set n) (_∩_ : L → L → L ) : Set (suc n) where - field - GFilter : F-Filter L PL _⊆_ _∩_ - Intersection : (D : F-Dense L PL _⊆_ _∩_ ) → { x : L } → F-Dense.dense D x → L - Generic : (D : F-Dense L PL _⊆_ _∩_ ) → { x : L } → ( y : F-Dense.dense D x) → F-Filter.filter GFilter (Intersection D y ) -
--- a/generic-filter.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,263 +0,0 @@ -open import Level -open import Ordinals -module generic-filter {n : Level } (O : Ordinals {n}) where - -import filter -open import zf -open import logic --- open import partfunc {n} O -import OD - -open import Relation.Nullary -open import Relation.Binary -open import Data.Empty -open import Relation.Binary -open import Relation.Binary.Core -open import Relation.Binary.PropositionalEquality -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -import BAlgbra - -open BAlgbra O - -open inOrdinal O -open OD O -open OD.OD -open ODAxiom odAxiom -import OrdUtil -import ODUtil -open Ordinals.Ordinals O -open Ordinals.IsOrdinals isOrdinal -open Ordinals.IsNext isNext -open OrdUtil O -open ODUtil O - - -import ODC - -open filter O - -open _∧_ -open _∨_ -open Bool - - -open HOD - -------- --- the set of finite partial functions from ω to 2 --- --- - -open import Data.List hiding (filter) -open import Data.Maybe - -import OPair -open OPair O - -ODSuc : (y : HOD) → infinite ∋ y → HOD -ODSuc y lt = Union (y , (y , y)) - -data Hω2 : (i : Nat) ( x : Ordinal ) → Set n where - hφ : Hω2 0 o∅ - h0 : {i : Nat} {x : Ordinal } → Hω2 i x → - Hω2 (Suc i) (& (Union ((< nat→ω i , nat→ω 0 >) , * x ))) - h1 : {i : Nat} {x : Ordinal } → Hω2 i x → - Hω2 (Suc i) (& (Union ((< nat→ω i , nat→ω 1 >) , * x ))) - he : {i : Nat} {x : Ordinal } → Hω2 i x → - Hω2 (Suc i) x - -record Hω2r (x : Ordinal) : Set n where - field - count : Nat - hω2 : Hω2 count x - -open Hω2r - -HODω2 : HOD -HODω2 = record { od = record { def = λ x → Hω2r x } ; odmax = next o∅ ; <odmax = odmax0 } where - P : (i j : Nat) (x : Ordinal ) → HOD - P i j x = ((nat→ω i , nat→ω i) , (nat→ω i , nat→ω j)) , * x - nat1 : (i : Nat) (x : Ordinal) → & (nat→ω i) o< next x - nat1 i x = next< nexto∅ ( <odmax infinite (ω∋nat→ω {i})) - lemma1 : (i j : Nat) (x : Ordinal ) → osuc (& (P i j x)) o< next x - lemma1 i j x = osuc<nx (pair-<xy (pair-<xy (pair-<xy (nat1 i x) (nat1 i x) ) (pair-<xy (nat1 i x) (nat1 j x) ) ) - (subst (λ k → k o< next x) (sym &iso) x<nx )) - lemma : (i j : Nat) (x : Ordinal ) → & (Union (P i j x)) o< next x - lemma i j x = next< (lemma1 i j x ) ho< - odmax0 : {y : Ordinal} → Hω2r y → y o< next o∅ - odmax0 {y} r with hω2 r - ... | hφ = x<nx - ... | h0 {i} {x} t = next< (odmax0 record { count = i ; hω2 = t }) (lemma i 0 x) - ... | h1 {i} {x} t = next< (odmax0 record { count = i ; hω2 = t }) (lemma i 1 x) - ... | he {i} {x} t = next< (odmax0 record { count = i ; hω2 = t }) x<nx - -3→Hω2 : List (Maybe Two) → HOD -3→Hω2 t = list→hod t 0 where - list→hod : List (Maybe Two) → Nat → HOD - list→hod [] _ = od∅ - list→hod (just i0 ∷ t) i = Union (< nat→ω i , nat→ω 0 > , ( list→hod t (Suc i) )) - list→hod (just i1 ∷ t) i = Union (< nat→ω i , nat→ω 1 > , ( list→hod t (Suc i) )) - list→hod (nothing ∷ t) i = list→hod t (Suc i ) - -Hω2→3 : (x : HOD) → HODω2 ∋ x → List (Maybe Two) -Hω2→3 x = lemma where - lemma : { y : Ordinal } → Hω2r y → List (Maybe Two) - lemma record { count = 0 ; hω2 = hφ } = [] - lemma record { count = (Suc i) ; hω2 = (h0 hω3) } = just i0 ∷ lemma record { count = i ; hω2 = hω3 } - lemma record { count = (Suc i) ; hω2 = (h1 hω3) } = just i1 ∷ lemma record { count = i ; hω2 = hω3 } - lemma record { count = (Suc i) ; hω2 = (he hω3) } = nothing ∷ lemma record { count = i ; hω2 = hω3 } - -ω→2 : HOD -ω→2 = Power infinite - -ω→2f : (x : HOD) → ω→2 ∋ x → Nat → Two -ω→2f x lt n with ODC.∋-p O x (nat→ω n) -ω→2f x lt n | yes p = i1 -ω→2f x lt n | no ¬p = i0 - -fω→2-sel : ( f : Nat → Two ) (x : HOD) → Set n -fω→2-sel f x = (infinite ∋ x) ∧ ( (lt : odef infinite (& x) ) → f (ω→nat x lt) ≡ i1 ) - -fω→2 : (Nat → Two) → HOD -fω→2 f = Select infinite (fω→2-sel f) - -open _==_ - -import Axiom.Extensionality.Propositional -postulate f-extensionality : { n m : Level} → Axiom.Extensionality.Propositional.Extensionality n m - -ω2∋f : (f : Nat → Two) → ω→2 ∋ fω→2 f -ω2∋f f = power← infinite (fω→2 f) (λ {x} lt → proj1 ((proj2 (selection {fω→2-sel f} {infinite} )) lt)) - -ω→2f≡i1 : (X i : HOD) → (iω : infinite ∋ i) → (lt : ω→2 ∋ X ) → ω→2f X lt (ω→nat i iω) ≡ i1 → X ∋ i -ω→2f≡i1 X i iω lt eq with ODC.∋-p O X (nat→ω (ω→nat i iω)) -ω→2f≡i1 X i iω lt eq | yes p = subst (λ k → X ∋ k ) (nat→ω-iso iω) p - -open _⊆_ - --- someday ... --- postulate --- ω→2f-iso : (X : HOD) → ( lt : ω→2 ∋ X ) → fω→2 ( ω→2f X lt ) =h= X --- fω→2-iso : (f : Nat → Two) → ω→2f ( fω→2 f ) (ω2∋f f) ≡ f - -record CountableOrdinal : Set (suc (suc n)) where - field - ctl→ : Nat → Ordinal - ctl← : Ordinal → Nat - ctl-iso→ : { x : Ordinal } → ctl→ (ctl← x ) ≡ x - ctl-iso← : { x : Nat } → ctl← (ctl→ x ) ≡ x - -record CountableHOD : Set (suc (suc n)) where - field - mhod : HOD - mtl→ : Nat → Ordinal - mtl→∈P : (i : Nat) → odef mhod (mtl→ i) - mtl← : (x : Ordinal) → odef mhod x → Nat - mtl-iso→ : { x : Ordinal } → (lt : odef mhod x ) → mtl→ (mtl← x lt ) ≡ x - mtl-iso← : { x : Nat } → mtl← (mtl→ x ) (mtl→∈P x) ≡ x - - -open CountableOrdinal -open CountableHOD - ----- --- a(n) ∈ M --- ∃ q ∈ Power P → q ∈ a(n) ∧ p(n) ⊆ q --- -PGHOD : (i : Nat) → (C : CountableOrdinal) → (P : HOD) → (p : Ordinal) → HOD -PGHOD i C P p = record { od = record { def = λ x → - odef (Power P) x ∧ odef (* (ctl→ C i)) x ∧ ( (y : Ordinal ) → odef (* p) y → odef (* x) y ) } - ; odmax = odmax (Power P) ; <odmax = λ {y} lt → <odmax (Power P) (proj1 lt) } - ---- --- p(n+1) = if PGHOD n qn otherwise p(n) --- -next-p : (C : CountableOrdinal) (P : HOD ) (i : Nat) → (p : Ordinal) → Ordinal -next-p C P i p with is-o∅ ( & (PGHOD i C P p)) -next-p C P i p | yes y = p -next-p C P i p | no not = & (ODC.minimal O (PGHOD i C P p ) (λ eq → not (=od∅→≡o∅ eq))) -- axiom of choice - ---- --- ascendant search on p(n) --- -find-p : (C : CountableOrdinal) (P : HOD ) (i : Nat) → (x : Ordinal) → Ordinal -find-p C P Zero x = x -find-p C P (Suc i) x = find-p C P i ( next-p C P i x ) - ---- --- G = { r ∈ Power P | ∃ n → r ⊆ p(n) } --- -record PDN (C : CountableOrdinal) (P : HOD ) (x : Ordinal) : Set n where - field - gr : Nat - pn<gr : (y : Ordinal) → odef (* x) y → odef (* (find-p C P gr o∅)) y - x∈PP : odef (Power P) x - -open PDN - ---- --- G as a HOD --- -PDHOD : (C : CountableOrdinal) → (P : HOD ) → HOD -PDHOD C P = record { od = record { def = λ x → PDN C P x } - ; odmax = odmax (Power P) ; <odmax = λ {y} lt → <odmax (Power P) {y} (PDN.x∈PP lt) } - -open PDN - ----- --- Generic Filter on Power P for HOD's Countable Ordinal (G ⊆ Power P ≡ G i.e. Nat → P → Set ) --- --- p 0 ≡ ∅ --- p (suc n) = if ∃ q ∈ * ( ctl→ n ) ∧ p n ⊆ q → q (axiom of choice) ---- else p n - -P∅ : {P : HOD} → odef (Power P) o∅ -P∅ {P} = subst (λ k → odef (Power P) k ) ord-od∅ (lemma o∅ o∅≡od∅) where - lemma : (x : Ordinal ) → * x ≡ od∅ → odef (Power P) (& od∅) - lemma x eq = power← P od∅ (λ {x} lt → ⊥-elim (¬x<0 lt )) -x<y→∋ : {x y : Ordinal} → odef (* x) y → * x ∋ * y -x<y→∋ {x} {y} lt = subst (λ k → odef (* x) k ) (sym &iso) lt - -P-GenericFilter : (C : CountableOrdinal) → (P : HOD ) → GenericFilter P -P-GenericFilter C P = record { - genf = record { filter = PDHOD C P ; f⊆PL = f⊆PL ; filter1 = {!!} ; filter2 = {!!} } - ; generic = λ D → {!!} - } where - find-p-⊆P : (C : CountableOrdinal) (P : HOD ) (i : Nat) → (x y : Ordinal) → odef (Power P) x → odef (* (find-p C P i x)) y → odef P y - find-p-⊆P C P Zero x y Px Py = subst (λ k → odef P k ) &iso - ( incl (ODC.power→⊆ O P (* x) (d→∋ (Power P) Px)) (x<y→∋ Py)) - find-p-⊆P C P (Suc i) x y Px Py = find-p-⊆P C P i (next-p C P i x) y {!!} {!!} - f⊆PL : PDHOD C P ⊆ Power P - f⊆PL = record { incl = λ {x} lt → power← P x (λ {y} y<x → - find-p-⊆P C P (gr lt) o∅ (& y) P∅ (pn<gr lt (& y) (subst (λ k → odef k (& y)) (sym *iso) y<x))) } - - -open GenericFilter -open Filter - -record Incompatible (P : HOD ) : Set (suc (suc n)) where - field - except : HOD → ( HOD ∧ HOD ) - incompatible : { p : HOD } → Power P ∋ p → Power P ∋ proj1 (except p ) → Power P ∋ proj2 (except p ) - → ( p ⊆ proj1 (except p) ) ∧ ( p ⊆ proj2 (except p) ) - → ∀ ( r : HOD ) → Power P ∋ r → ¬ (( proj1 (except p) ⊆ r ) ∧ ( proj2 (except p) ⊆ r )) - -lemma725 : (M : CountableHOD ) (C : CountableOrdinal) (P : HOD ) → mhod M ∋ Power P - → Incompatible P → ¬ ( mhod M ∋ filter ( genf ( P-GenericFilter C P ))) -lemma725 = {!!} - -lemma725-1 : Incompatible HODω2 -lemma725-1 = {!!} - -lemma726 : (C : CountableOrdinal) (P : HOD ) - → Union ( filter ( genf ( P-GenericFilter C HODω2 ))) =h= ω→2 -lemma726 = {!!} - --- --- val x G = { val y G | ∃ p → G ∋ p → x ∋ < y , p > } --- --- W (ω , H ( ω , 2 )) = { p ∈ ( Nat → H (ω , 2) ) | { i ∈ Nat → p i ≠ i1 } is finite } --- - - -
--- a/logic.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,57 +0,0 @@ -module logic where - -open import Level -open import Relation.Nullary -open import Relation.Binary hiding (_⇔_ ) -open import Data.Empty - -data One {n : Level } : Set n where - OneObj : One - -data Two : Set where - i0 : Two - i1 : Two - -data Bool : Set where - true : Bool - false : Bool - -record _∧_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where - constructor ⟪_,_⟫ - field - proj1 : A - proj2 : B - -data _∨_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where - case1 : A → A ∨ B - case2 : B → A ∨ B - -_⇔_ : {n m : Level } → ( A : Set n ) ( B : Set m ) → Set (n ⊔ m) -_⇔_ A B = ( A → B ) ∧ ( B → A ) - -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 ) - -double-neg : {n : Level } {A : Set n} → A → ¬ ¬ A -double-neg A notnot = notnot A - -double-neg2 : {n : Level } {A : Set n} → ¬ ¬ ¬ A → ¬ A -double-neg2 notnot A = notnot ( double-neg A ) - -de-morgan : {n : Level } {A B : Set n} → A ∧ B → ¬ ( (¬ A ) ∨ (¬ B ) ) -de-morgan {n} {A} {B} and (case1 ¬A) = ⊥-elim ( ¬A ( _∧_.proj1 and )) -de-morgan {n} {A} {B} and (case2 ¬B) = ⊥-elim ( ¬B ( _∧_.proj2 and )) - -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 - -dont-orb : {n m : Level} {A : Set n} { B : Set m } → A ∨ B → ¬ B → A -dont-orb {A} {B} (case2 b) ¬B = ⊥-elim ( ¬B b ) -dont-orb {A} {B} (case1 a) ¬B = a - - -infixr 130 _∧_ -infixr 140 _∨_ -infixr 150 _⇔_ -
--- a/nat.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,46 +0,0 @@ -module nat where - -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -open import Data.Empty -open import Relation.Nullary -open import Relation.Binary.PropositionalEquality -open import logic - - -nat-<> : { x y : Nat } → x < y → y < x → ⊥ -nat-<> (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x - -nat-<≡ : { x : Nat } → x < x → ⊥ -nat-<≡ (s≤s lt) = nat-<≡ lt - -nat-≡< : { x y : Nat } → x ≡ y → x < y → ⊥ -nat-≡< refl lt = nat-<≡ lt - -¬a≤a : {la : Nat} → Suc la ≤ la → ⊥ -¬a≤a (s≤s lt) = ¬a≤a lt - -a<sa : {la : Nat} → la < Suc la -a<sa {Zero} = s≤s z≤n -a<sa {Suc la} = s≤s a<sa - -=→¬< : {x : Nat } → ¬ ( x < x ) -=→¬< {Zero} () -=→¬< {Suc x} (s≤s lt) = =→¬< lt - ->→¬< : {x y : Nat } → (x < y ) → ¬ ( y < x ) ->→¬< (s≤s x<y) (s≤s y<x) = >→¬< x<y y<x - -<-∨ : { x y : Nat } → x < Suc y → ( (x ≡ y ) ∨ (x < y) ) -<-∨ {Zero} {Zero} (s≤s z≤n) = case1 refl -<-∨ {Zero} {Suc y} (s≤s lt) = case2 (s≤s z≤n) -<-∨ {Suc x} {Zero} (s≤s ()) -<-∨ {Suc x} {Suc y} (s≤s lt) with <-∨ {x} {y} lt -<-∨ {Suc x} {Suc y} (s≤s lt) | case1 eq = case1 (cong (λ k → Suc k ) eq) -<-∨ {Suc x} {Suc y} (s≤s lt) | case2 lt1 = case2 (s≤s lt1) - -max : (x y : Nat) → Nat -max Zero Zero = Zero -max Zero (Suc x) = (Suc x) -max (Suc x) Zero = (Suc x) -max (Suc x) (Suc y) = Suc ( max x y ) -
--- a/ordinal.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,282 +0,0 @@ -open import Level -module ordinal where - -open import logic -open import nat -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -open import Data.Empty -open import Relation.Binary.PropositionalEquality -open import Relation.Binary.Definitions -open import Data.Nat.Properties -open import Relation.Nullary -open import Relation.Binary.Core - ----- --- --- Countable Ordinals --- - -data OrdinalD {n : Level} : (lv : Nat) → Set n where - Φ : (lv : Nat) → OrdinalD lv - OSuc : (lv : Nat) → OrdinalD {n} lv → OrdinalD lv - -record Ordinal {n : Level} : Set n where - constructor ordinal - field - lv : Nat - ord : OrdinalD {n} lv - -data _d<_ {n : Level} : {lx ly : Nat} → OrdinalD {n} lx → OrdinalD {n} ly → Set n where - Φ< : {lx : Nat} → {x : OrdinalD {n} lx} → Φ lx d< OSuc lx x - s< : {lx : Nat} → {x y : OrdinalD {n} lx} → x d< y → OSuc lx x d< OSuc lx y - -open Ordinal - -_o<_ : {n : Level} ( x y : Ordinal ) → Set n -_o<_ x y = (lv x < lv y ) ∨ ( ord x d< ord y ) - -s<refl : {n : Level } {lx : Nat } { x : OrdinalD {n} lx } → x d< OSuc lx x -s<refl {n} {lv} {Φ lv} = Φ< -s<refl {n} {lv} {OSuc lv x} = s< s<refl - -trio<> : {n : Level} → {lx : Nat} {x : OrdinalD {n} lx } { y : OrdinalD lx } → y d< x → x d< y → ⊥ -trio<> {n} {lx} {.(OSuc lx _)} {.(OSuc lx _)} (s< s) (s< t) = trio<> s t -trio<> {n} {lx} {.(OSuc lx _)} {.(Φ lx)} Φ< () - -d<→lv : {n : Level} {x y : Ordinal {n}} → ord x d< ord y → lv x ≡ lv y -d<→lv Φ< = refl -d<→lv (s< lt) = refl - -o∅ : {n : Level} → Ordinal {n} -o∅ = record { lv = Zero ; ord = Φ Zero } - -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 - -≡→¬d< : {n : Level} → {lv : Nat} → {x : OrdinalD {n} lv } → x d< x → ⊥ -≡→¬d< {n} {lx} {OSuc lx y} (s< t) = ≡→¬d< t - -trio<≡ : {n : Level} → {lx : Nat} {x : OrdinalD {n} lx } { y : OrdinalD lx } → x ≡ y → x d< y → ⊥ -trio<≡ refl = ≡→¬d< - -trio>≡ : {n : Level} → {lx : Nat} {x : OrdinalD {n} lx } { y : OrdinalD lx } → x ≡ y → y d< x → ⊥ -trio>≡ refl = ≡→¬d< - -triOrdd : {n : Level} → {lx : Nat} → Trichotomous _≡_ ( _d<_ {n} {lx} {lx} ) -triOrdd {_} {lv} (Φ lv) (Φ lv) = tri≈ ≡→¬d< refl ≡→¬d< -triOrdd {_} {lv} (Φ lv) (OSuc lv y) = tri< Φ< (λ ()) ( λ lt → trio<> lt Φ< ) -triOrdd {_} {lv} (OSuc lv x) (Φ lv) = tri> (λ lt → trio<> lt Φ<) (λ ()) Φ< -triOrdd {_} {lv} (OSuc lv x) (OSuc lv y) with triOrdd x y -triOrdd {_} {lv} (OSuc lv x) (OSuc lv y) | tri< a ¬b ¬c = tri< (s< a) (λ tx=ty → trio<≡ tx=ty (s< a) ) ( λ lt → trio<> lt (s< a) ) -triOrdd {_} {lv} (OSuc lv x) (OSuc lv x) | tri≈ ¬a refl ¬c = tri≈ ≡→¬d< refl ≡→¬d< -triOrdd {_} {lv} (OSuc lv x) (OSuc lv y) | tri> ¬a ¬b c = tri> ( λ lt → trio<> lt (s< c) ) (λ tx=ty → trio>≡ tx=ty (s< c) ) (s< c) - -osuc : {n : Level} ( x : Ordinal {n} ) → Ordinal {n} -osuc record { lv = lx ; ord = ox } = record { lv = lx ; ord = OSuc lx ox } - -<-osuc : {n : Level} { x : Ordinal {n} } → x o< osuc x -<-osuc {n} {record { lv = lx ; ord = Φ .lx }} = case2 Φ< -<-osuc {n} {record { lv = lx ; ord = OSuc .lx ox }} = case2 ( s< s<refl ) - -o<¬≡ : {n : Level } { ox oy : Ordinal {suc n}} → ox ≡ oy → ox o< oy → ⊥ -o<¬≡ {_} {ox} {ox} refl (case1 lt) = =→¬< lt -o<¬≡ {_} {ox} {ox} refl (case2 (s< lt)) = trio<≡ refl lt - -¬x<0 : {n : Level} → { x : Ordinal {suc n} } → ¬ ( x o< o∅ {suc n} ) -¬x<0 {n} {x} (case1 ()) -¬x<0 {n} {x} (case2 ()) - -o<> : {n : Level} → {x y : Ordinal {n} } → y o< x → x o< y → ⊥ -o<> {n} {x} {y} (case1 x₁) (case1 x₂) = nat-<> x₁ x₂ -o<> {n} {x} {y} (case1 x₁) (case2 x₂) = nat-≡< (sym (d<→lv x₂)) x₁ -o<> {n} {x} {y} (case2 x₁) (case1 x₂) = nat-≡< (sym (d<→lv x₁)) x₂ -o<> {n} {record { lv = lv₁ ; ord = .(OSuc lv₁ _) }} {record { lv = .lv₁ ; ord = .(Φ lv₁) }} (case2 Φ<) (case2 ()) -o<> {n} {record { lv = lv₁ ; ord = .(OSuc lv₁ _) }} {record { lv = .lv₁ ; ord = .(OSuc lv₁ _) }} (case2 (s< y<x)) (case2 (s< x<y)) = - o<> (case2 y<x) (case2 x<y) - -orddtrans : {n : Level} {lx : Nat} {x y z : OrdinalD {n} lx } → x d< y → y d< z → x d< z -orddtrans {_} {lx} {.(Φ lx)} {.(OSuc lx _)} {.(OSuc lx _)} Φ< (s< y<z) = Φ< -orddtrans {_} {lx} {.(OSuc lx _)} {.(OSuc lx _)} {.(OSuc lx _)} (s< x<y) (s< y<z) = s< ( orddtrans x<y y<z ) - -osuc-≡< : {n : Level} { a x : Ordinal {n} } → x o< osuc a → (x ≡ a ) ∨ (x o< a) -osuc-≡< {n} {a} {x} (case1 lt) = case2 (case1 lt) -osuc-≡< {n} {record { lv = lv₁ ; ord = Φ .lv₁ }} {record { lv = .lv₁ ; ord = .(Φ lv₁) }} (case2 Φ<) = case1 refl -osuc-≡< {n} {record { lv = lv₁ ; ord = OSuc .lv₁ ord₁ }} {record { lv = .lv₁ ; ord = .(Φ lv₁) }} (case2 Φ<) = case2 (case2 Φ<) -osuc-≡< {n} {record { lv = lv₁ ; ord = Φ .lv₁ }} {record { lv = .lv₁ ; ord = .(OSuc lv₁ _) }} (case2 (s< ())) -osuc-≡< {n} {record { lv = la ; ord = OSuc la oa }} {record { lv = la ; ord = (OSuc la ox) }} (case2 (s< lt)) with - osuc-≡< {n} {record { lv = la ; ord = oa }} {record { lv = la ; ord = ox }} (case2 lt ) -... | case1 refl = case1 refl -... | case2 (case2 x) = case2 (case2( s< x) ) -... | case2 (case1 x) = ⊥-elim (¬a≤a x) - -osuc-< : {n : Level} { x y : Ordinal {n} } → y o< osuc x → x o< y → ⊥ -osuc-< {n} {x} {y} y<ox x<y with osuc-≡< y<ox -osuc-< {n} {x} {x} y<ox (case1 x₁) | case1 refl = ⊥-elim (¬a≤a x₁) -osuc-< {n} {x} {x} (case1 x₂) (case2 x₁) | case1 refl = ⊥-elim (¬a≤a x₂) -osuc-< {n} {x} {x} (case2 x₂) (case2 x₁) | case1 refl = ≡→¬d< x₁ -osuc-< {n} {x} {y} y<ox (case1 x₂) | case2 (case1 x₁) = nat-<> x₂ x₁ -osuc-< {n} {x} {y} y<ox (case1 x₂) | case2 (case2 x₁) = nat-≡< (sym (d<→lv x₁)) x₂ -osuc-< {n} {x} {y} y<ox (case2 x<y) | case2 y<x = o<> (case2 x<y) y<x - - -ordtrans : {n : Level} {x y z : Ordinal {n} } → x o< y → y o< z → x o< z -ordtrans {n} {x} {y} {z} (case1 x₁) (case1 x₂) = case1 ( <-trans x₁ x₂ ) -ordtrans {n} {x} {y} {z} (case1 x₁) (case2 x₂) with d<→lv x₂ -... | refl = case1 x₁ -ordtrans {n} {x} {y} {z} (case2 x₁) (case1 x₂) with d<→lv x₁ -... | refl = case1 x₂ -ordtrans {n} {x} {y} {z} (case2 x₁) (case2 x₂) with d<→lv x₁ | d<→lv x₂ -... | refl | refl = case2 ( orddtrans x₁ x₂ ) - -trio< : {n : Level } → Trichotomous {suc n} _≡_ _o<_ -trio< a b with <-cmp (lv a) (lv b) -trio< a b | tri< a₁ ¬b ¬c = tri< (case1 a₁) (λ refl → ¬b (cong ( λ x → lv x ) refl ) ) lemma1 where - lemma1 : ¬ (Suc (lv b) ≤ lv a) ∨ (ord b d< ord a) - lemma1 (case1 x) = ¬c x - lemma1 (case2 x) = ⊥-elim (nat-≡< (sym ( d<→lv x )) a₁ ) -trio< a b | tri> ¬a ¬b c = tri> lemma1 (λ refl → ¬b (cong ( λ x → lv x ) refl ) ) (case1 c) where - lemma1 : ¬ (Suc (lv a) ≤ lv b) ∨ (ord a d< ord b) - lemma1 (case1 x) = ¬a x - lemma1 (case2 x) = ⊥-elim (nat-≡< (sym ( d<→lv x )) c ) -trio< a b | tri≈ ¬a refl ¬c with triOrdd ( ord a ) ( ord b ) -trio< record { lv = .(lv b) ; ord = x } b | tri≈ ¬a refl ¬c | tri< a ¬b ¬c₁ = tri< (case2 a) (λ refl → ¬b (lemma1 refl )) lemma2 where - lemma1 : (record { lv = _ ; ord = x }) ≡ b → x ≡ ord b - lemma1 refl = refl - lemma2 : ¬ (Suc (lv b) ≤ lv b) ∨ (ord b d< x) - lemma2 (case1 x) = ¬a x - lemma2 (case2 x) = trio<> x a -trio< record { lv = .(lv b) ; ord = x } b | tri≈ ¬a refl ¬c | tri> ¬a₁ ¬b c = tri> lemma2 (λ refl → ¬b (lemma1 refl )) (case2 c) where - lemma1 : (record { lv = _ ; ord = x }) ≡ b → x ≡ ord b - lemma1 refl = refl - lemma2 : ¬ (Suc (lv b) ≤ lv b) ∨ (x d< ord b) - lemma2 (case1 x) = ¬a x - lemma2 (case2 x) = trio<> x c -trio< record { lv = .(lv b) ; ord = x } b | tri≈ ¬a refl ¬c | tri≈ ¬a₁ refl ¬c₁ = tri≈ lemma1 refl lemma1 where - lemma1 : ¬ (Suc (lv b) ≤ lv b) ∨ (ord b d< ord b) - lemma1 (case1 x) = ¬a x - lemma1 (case2 x) = ≡→¬d< x - - -open _∧_ - -TransFinite : {n m : Level} → { ψ : Ordinal {suc n} → Set m } - → ( ∀ (lx : Nat ) → ( (x : Ordinal {suc n} ) → x o< ordinal lx (Φ lx) → ψ x ) → ψ ( record { lv = lx ; ord = Φ lx } ) ) - → ( ∀ (lx : Nat ) → (x : OrdinalD lx ) → ( (y : Ordinal {suc n} ) → y o< ordinal lx (OSuc lx x) → ψ y ) → ψ ( record { lv = lx ; ord = OSuc lx x } ) ) - → ∀ (x : Ordinal) → ψ x -TransFinite {n} {m} {ψ} caseΦ caseOSuc x = proj1 (TransFinite1 (lv x) (ord x) ) where - TransFinite1 : (lx : Nat) (ox : OrdinalD lx ) → ψ (ordinal lx ox) ∧ ( ( (x : Ordinal {suc n} ) → x o< ordinal lx ox → ψ x ) ) - TransFinite1 Zero (Φ 0) = ⟪ caseΦ Zero lemma , lemma1 ⟫ where - lemma : (x : Ordinal) → x o< ordinal Zero (Φ Zero) → ψ x - lemma x (case1 ()) - lemma x (case2 ()) - lemma1 : (x : Ordinal) → x o< ordinal Zero (Φ Zero) → ψ x - lemma1 x (case1 ()) - lemma1 x (case2 ()) - TransFinite1 (Suc lx) (Φ (Suc lx)) = ⟪ caseΦ (Suc lx) (λ x → lemma (lv x) (ord x)) , (λ x → lemma (lv x) (ord x)) ⟫ where - lemma0 : (ly : Nat) (oy : OrdinalD ly ) → ordinal ly oy o< ordinal lx (Φ lx) → ψ (ordinal ly oy) - lemma0 ly oy lt = proj2 ( TransFinite1 lx (Φ lx) ) (ordinal ly oy) lt - lemma : (ly : Nat) (oy : OrdinalD ly ) → ordinal ly oy o< ordinal (Suc lx) (Φ (Suc lx)) → ψ (ordinal ly oy) - lemma lx1 ox1 (case1 lt) with <-∨ lt - lemma lx (Φ lx) (case1 lt) | case1 refl = proj1 ( TransFinite1 lx (Φ lx) ) - lemma lx (Φ lx) (case1 lt) | case2 lt1 = lemma0 lx (Φ lx) (case1 lt1) - lemma lx (OSuc lx ox1) (case1 lt) | case1 refl = caseOSuc lx ox1 lemma2 where - lemma2 : (y : Ordinal) → (Suc (lv y) ≤ lx) ∨ (ord y d< OSuc lx ox1) → ψ y - lemma2 y lt1 with osuc-≡< lt1 - lemma2 y lt1 | case1 refl = lemma lx ox1 (case1 a<sa) - lemma2 y lt1 | case2 t = proj2 (TransFinite1 lx ox1) y t - lemma lx1 (OSuc lx1 ox1) (case1 lt) | case2 lt1 = caseOSuc lx1 ox1 lemma2 where - lemma2 : (y : Ordinal) → (Suc (lv y) ≤ lx1) ∨ (ord y d< OSuc lx1 ox1) → ψ y - lemma2 y lt2 with osuc-≡< lt2 - lemma2 y lt2 | case1 refl = lemma lx1 ox1 (ordtrans lt2 (case1 lt)) - lemma2 y lt2 | case2 (case1 lt3) = proj2 (TransFinite1 lx (Φ lx)) y (case1 (<-trans lt3 lt1 )) - lemma2 y lt2 | case2 (case2 lt3) with d<→lv lt3 - ... | refl = proj2 (TransFinite1 lx (Φ lx)) y (case1 lt1) - TransFinite1 lx (OSuc lx ox) = ⟪ caseOSuc lx ox lemma , lemma ⟫ where - lemma : (y : Ordinal) → y o< ordinal lx (OSuc lx ox) → ψ y - lemma y lt with osuc-≡< lt - lemma y lt | case1 refl = proj1 ( TransFinite1 lx ox ) - lemma y lt | case2 lt1 = proj2 ( TransFinite1 lx ox ) y lt1 - --- record CountableOrdinal {n : Level} : Set (suc (suc n)) where --- field --- ctl→ : Nat → Ordinal {suc n} --- ctl← : Ordinal → Nat --- ctl-iso→ : { x : Ordinal } → ctl→ (ctl← x ) ≡ x --- ctl-iso← : { x : Nat } → ctl← (ctl→ x ) ≡ x --- --- is-C-Ordinal : {n : Level} → CountableOrdinal {n} --- is-C-Ordinal {n} = record { --- ctl→ = {!!} --- ; ctl← = λ x → TransFinite {n} (λ lx lt → Zero ) ctl01 x --- ; ctl-iso→ = {!!} --- ; ctl-iso← = {!!} --- } where --- ctl01 : (lx : Nat) (x : OrdinalD lx) → ((y : Ordinal) → y o< ordinal lx (OSuc lx x) → Nat) → Nat --- ctl01 Zero (Φ Zero) prev = Zero --- ctl01 Zero (OSuc Zero x) prev = Suc ( prev (ordinal Zero x) (ordtrans <-osuc <-osuc )) --- ctl01 (Suc lx) (Φ (Suc lx)) prev = Suc ( prev (ordinal lx {!!}) {!!}) --- ctl01 (Suc lx) (OSuc (Suc lx) x) prev = Suc ( prev (ordinal (Suc lx) x) (ordtrans <-osuc <-osuc )) - -open import Ordinals - -C-Ordinal : {n : Level} → Ordinals {suc n} -C-Ordinal {n} = record { - Ordinal = Ordinal {suc n} - ; o∅ = o∅ - ; osuc = osuc - ; _o<_ = _o<_ - ; next = next - ; isOrdinal = record { - ordtrans = ordtrans - ; trio< = trio< - ; ¬x<0 = ¬x<0 - ; <-osuc = <-osuc - ; osuc-≡< = osuc-≡< - ; TransFinite = TransFinite2 - ; Oprev-p = Oprev-p - } ; - isNext = record { - x<nx = x<nx - ; osuc<nx = λ {x} {y} → osuc<nx {x} {y} - ; ¬nx<nx = ¬nx<nx - } - } where - next : Ordinal {suc n} → Ordinal {suc n} - next (ordinal lv ord) = ordinal (Suc lv) (Φ (Suc lv)) - x<nx : { y : Ordinal } → (y o< next y ) - x<nx = case1 a<sa - osuc<nx : { x y : Ordinal } → x o< next y → osuc x o< next y - osuc<nx (case1 lt) = case1 lt - ¬nx<nx : {x y : Ordinal} → y o< x → x o< next y → ¬ ((z : Ordinal) → ¬ (x ≡ osuc z)) - ¬nx<nx {x} {y} = lemma2 x where - lemma2 : (x : Ordinal) → y o< x → x o< next y → ¬ ((z : Ordinal) → ¬ x ≡ osuc z) - lemma2 (ordinal Zero (Φ 0)) (case2 ()) (case1 (s≤s z≤n)) not - lemma2 (ordinal Zero (OSuc 0 dx)) (case2 Φ<) (case1 (s≤s z≤n)) not = not _ refl - lemma2 (ordinal Zero (OSuc 0 dx)) (case2 (s< x)) (case1 (s≤s z≤n)) not = not _ refl - lemma2 (ordinal (Suc lx) (OSuc (Suc lx) ox)) y<x (case1 (s≤s (s≤s lt))) not = not _ refl - lemma2 (ordinal (Suc lx) (Φ (Suc lx))) (case1 x) (case1 (s≤s (s≤s lt))) not = lemma3 x lt where - lemma3 : {n l : Nat} → (Suc (Suc n) ≤ Suc l) → l ≤ n → ⊥ - lemma3 (s≤s sn≤l) (s≤s l≤n) = lemma3 sn≤l l≤n - open Oprev - Oprev-p : (x : Ordinal) → Dec ( Oprev (Ordinal {suc n}) osuc x ) - Oprev-p (ordinal lv (Φ lv)) = no (λ not → lemma (oprev not) (oprev=x not) ) where - lemma : (x : Ordinal) → osuc x ≡ (ordinal lv (Φ lv)) → ⊥ - lemma x () - Oprev-p (ordinal lv (OSuc lv ox)) = yes record { oprev = ordinal lv ox ; oprev=x = refl } - ord1 : Set (suc n) - ord1 = Ordinal {suc n} - TransFinite2 : { ψ : ord1 → Set (suc (suc n)) } - → ( (x : ord1) → ( (y : ord1 ) → y o< x → ψ y ) → ψ x ) - → ∀ (x : ord1) → ψ x - TransFinite2 {ψ} lt x = TransFinite {n} {suc (suc n)} {ψ} caseΦ caseOSuc x where - caseΦ : (lx : Nat) → ((x₁ : Ordinal) → x₁ o< ordinal lx (Φ lx) → ψ x₁) → - ψ (record { lv = lx ; ord = Φ lx }) - caseΦ lx prev = lt (ordinal lx (Φ lx) ) prev - caseOSuc : (lx : Nat) (x₁ : OrdinalD lx) → ((y : Ordinal) → y o< ordinal lx (OSuc lx x₁) → ψ y) → - ψ (record { lv = lx ; ord = OSuc lx x₁ }) - caseOSuc lx ox prev = lt (ordinal lx (OSuc lx ox)) prev - -
--- a/partfunc.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,227 +0,0 @@ -{-# OPTIONS --allow-unsolved-metas #-} -open import Level -open import Relation.Nullary -open import Relation.Binary.PropositionalEquality -open import Ordinals -module partfunc {n : Level } (O : Ordinals {n}) where - -open import logic -open import Relation.Binary -open import Data.Empty -open import Data.List hiding (filter) -open import Data.Maybe -open import Relation.Binary -open import Relation.Binary.Core -open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) -open import filter O - -open _∧_ -open _∨_ -open Bool - ----- --- --- Partial Function without ZF --- - -record PFunc (Dom : Set n) (Cod : Set n) : Set (suc n) where - field - dom : Dom → Set n - pmap : (x : Dom ) → dom x → Cod - meq : {x : Dom } → { p q : dom x } → pmap x p ≡ pmap x q - ----- --- --- PFunc (Lift n Nat) Cod is equivalent to List (Maybe Cod) --- - -data Findp : {Cod : Set n} → List (Maybe Cod) → (x : Nat) → Set where - v0 : {Cod : Set n} → {f : List (Maybe Cod)} → ( v : Cod ) → Findp ( just v ∷ f ) Zero - vn : {Cod : Set n} → {f : List (Maybe Cod)} {d : Maybe Cod} → {x : Nat} → Findp f x → Findp (d ∷ f) (Suc x) - -open PFunc - -find : {Cod : Set n} → (f : List (Maybe Cod) ) → (x : Nat) → Findp f x → Cod -find (just v ∷ _) 0 (v0 v) = v -find (_ ∷ n) (Suc i) (vn p) = find n i p - -findpeq : {Cod : Set n} → (f : List (Maybe Cod)) → {x : Nat} {p q : Findp f x } → find f x p ≡ find f x q -findpeq n {0} {v0 _} {v0 _} = refl -findpeq [] {Suc x} {()} -findpeq (just x₁ ∷ n) {Suc x} {vn p} {vn q} = findpeq n {x} {p} {q} -findpeq (nothing ∷ n) {Suc x} {vn p} {vn q} = findpeq n {x} {p} {q} - -List→PFunc : {Cod : Set n} → List (Maybe Cod) → PFunc (Lift n Nat) Cod -List→PFunc fp = record { dom = λ x → Lift n (Findp fp (lower x)) - ; pmap = λ x y → find fp (lower x) (lower y) - ; meq = λ {x} {p} {q} → findpeq fp {lower x} {lower p} {lower q} - } ----- --- --- to List (Maybe Two) is a Latice --- - -_3⊆b_ : (f g : List (Maybe Two)) → Bool -[] 3⊆b [] = true -[] 3⊆b (nothing ∷ g) = [] 3⊆b g -[] 3⊆b (_ ∷ g) = true -(nothing ∷ f) 3⊆b [] = f 3⊆b [] -(nothing ∷ f) 3⊆b (_ ∷ g) = f 3⊆b g -(just i0 ∷ f) 3⊆b (just i0 ∷ g) = f 3⊆b g -(just i1 ∷ f) 3⊆b (just i1 ∷ g) = f 3⊆b g -_ 3⊆b _ = false - -_3⊆_ : (f g : List (Maybe Two)) → Set -f 3⊆ g = (f 3⊆b g) ≡ true - -_3∩_ : (f g : List (Maybe Two)) → List (Maybe Two) -[] 3∩ (nothing ∷ g) = nothing ∷ ([] 3∩ g) -[] 3∩ g = [] -(nothing ∷ f) 3∩ [] = nothing ∷ f 3∩ [] -f 3∩ [] = [] -(just i0 ∷ f) 3∩ (just i0 ∷ g) = just i0 ∷ ( f 3∩ g ) -(just i1 ∷ f) 3∩ (just i1 ∷ g) = just i1 ∷ ( f 3∩ g ) -(_ ∷ f) 3∩ (_ ∷ g) = nothing ∷ ( f 3∩ g ) - -3∩⊆f : { f g : List (Maybe Two) } → (f 3∩ g ) 3⊆ f -3∩⊆f {[]} {[]} = refl -3∩⊆f {[]} {just _ ∷ g} = refl -3∩⊆f {[]} {nothing ∷ g} = 3∩⊆f {[]} {g} -3∩⊆f {just _ ∷ f} {[]} = refl -3∩⊆f {nothing ∷ f} {[]} = 3∩⊆f {f} {[]} -3∩⊆f {just i0 ∷ f} {just i0 ∷ g} = 3∩⊆f {f} {g} -3∩⊆f {just i1 ∷ f} {just i1 ∷ g} = 3∩⊆f {f} {g} -3∩⊆f {just i0 ∷ f} {just i1 ∷ g} = 3∩⊆f {f} {g} -3∩⊆f {just i1 ∷ f} {just i0 ∷ g} = 3∩⊆f {f} {g} -3∩⊆f {nothing ∷ f} {just _ ∷ g} = 3∩⊆f {f} {g} -3∩⊆f {just i0 ∷ f} {nothing ∷ g} = 3∩⊆f {f} {g} -3∩⊆f {just i1 ∷ f} {nothing ∷ g} = 3∩⊆f {f} {g} -3∩⊆f {nothing ∷ f} {nothing ∷ g} = 3∩⊆f {f} {g} - -3∩sym : { f g : List (Maybe Two) } → (f 3∩ g ) ≡ (g 3∩ f ) -3∩sym {[]} {[]} = refl -3∩sym {[]} {just _ ∷ g} = refl -3∩sym {[]} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {[]} {g}) -3∩sym {just _ ∷ f} {[]} = refl -3∩sym {nothing ∷ f} {[]} = cong (λ k → nothing ∷ k) (3∩sym {f} {[]}) -3∩sym {just i0 ∷ f} {just i0 ∷ g} = cong (λ k → just i0 ∷ k) (3∩sym {f} {g}) -3∩sym {just i0 ∷ f} {just i1 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) -3∩sym {just i1 ∷ f} {just i0 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) -3∩sym {just i1 ∷ f} {just i1 ∷ g} = cong (λ k → just i1 ∷ k) (3∩sym {f} {g}) -3∩sym {just i0 ∷ f} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) -3∩sym {just i1 ∷ f} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) -3∩sym {nothing ∷ f} {just i0 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) -3∩sym {nothing ∷ f} {just i1 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) -3∩sym {nothing ∷ f} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) - -3⊆-[] : { h : List (Maybe Two) } → [] 3⊆ h -3⊆-[] {[]} = refl -3⊆-[] {just _ ∷ h} = refl -3⊆-[] {nothing ∷ h} = 3⊆-[] {h} - -3⊆trans : { f g h : List (Maybe Two) } → f 3⊆ g → g 3⊆ h → f 3⊆ h -3⊆trans {[]} {[]} {[]} f<g g<h = refl -3⊆trans {[]} {[]} {just _ ∷ h} f<g g<h = refl -3⊆trans {[]} {[]} {nothing ∷ h} f<g g<h = 3⊆trans {[]} {[]} {h} refl g<h -3⊆trans {[]} {nothing ∷ g} {[]} f<g g<h = refl -3⊆trans {[]} {just _ ∷ g} {just _ ∷ h} f<g g<h = refl -3⊆trans {[]} {nothing ∷ g} {just _ ∷ h} f<g g<h = refl -3⊆trans {[]} {nothing ∷ g} {nothing ∷ h} f<g g<h = 3⊆trans {[]} {g} {h} f<g g<h -3⊆trans {nothing ∷ f} {[]} {[]} f<g g<h = f<g -3⊆trans {nothing ∷ f} {[]} {just _ ∷ h} f<g g<h = 3⊆trans {f} {[]} {h} f<g (3⊆-[] {h}) -3⊆trans {nothing ∷ f} {[]} {nothing ∷ h} f<g g<h = 3⊆trans {f} {[]} {h} f<g g<h -3⊆trans {nothing ∷ f} {nothing ∷ g} {[]} f<g g<h = 3⊆trans {f} {g} {[]} f<g g<h -3⊆trans {nothing ∷ f} {nothing ∷ g} {just _ ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h -3⊆trans {nothing ∷ f} {nothing ∷ g} {nothing ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h -3⊆trans {[]} {just i0 ∷ g} {[]} f<g () -3⊆trans {[]} {just i1 ∷ g} {[]} f<g () -3⊆trans {[]} {just x ∷ g} {nothing ∷ h} f<g g<h = 3⊆-[] {h} -3⊆trans {just i0 ∷ f} {[]} {h} () g<h -3⊆trans {just i1 ∷ f} {[]} {h} () g<h -3⊆trans {just x ∷ f} {just i0 ∷ g} {[]} f<g () -3⊆trans {just x ∷ f} {just i1 ∷ g} {[]} f<g () -3⊆trans {just i0 ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h -3⊆trans {just i1 ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h -3⊆trans {just x ∷ f} {just i0 ∷ g} {nothing ∷ h} f<g () -3⊆trans {just x ∷ f} {just i1 ∷ g} {nothing ∷ h} f<g () -3⊆trans {just i0 ∷ f} {nothing ∷ g} {_} () g<h -3⊆trans {just i1 ∷ f} {nothing ∷ g} {_} () g<h -3⊆trans {nothing ∷ f} {just i0 ∷ g} {[]} f<g () -3⊆trans {nothing ∷ f} {just i1 ∷ g} {[]} f<g () -3⊆trans {nothing ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h -3⊆trans {nothing ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h -3⊆trans {nothing ∷ f} {just i0 ∷ g} {nothing ∷ h} f<g () -3⊆trans {nothing ∷ f} {just i1 ∷ g} {nothing ∷ h} f<g () - -3⊆∩f : { f g h : List (Maybe Two) } → f 3⊆ g → f 3⊆ h → f 3⊆ (g 3∩ h ) -3⊆∩f {[]} {[]} {[]} f<g f<h = refl -3⊆∩f {[]} {[]} {x ∷ h} f<g f<h = 3⊆-[] {[] 3∩ (x ∷ h)} -3⊆∩f {[]} {x ∷ g} {h} f<g f<h = 3⊆-[] {(x ∷ g) 3∩ h} -3⊆∩f {nothing ∷ f} {[]} {[]} f<g f<h = 3⊆∩f {f} {[]} {[]} f<g f<h -3⊆∩f {nothing ∷ f} {[]} {just _ ∷ h} f<g f<h = f<g -3⊆∩f {nothing ∷ f} {[]} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {[]} {h} f<g f<h -3⊆∩f {just i0 ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h -3⊆∩f {just i1 ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h -3⊆∩f {nothing ∷ f} {just _ ∷ g} {[]} f<g f<h = f<h -3⊆∩f {nothing ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h -3⊆∩f {nothing ∷ f} {just i0 ∷ g} {just i1 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h -3⊆∩f {nothing ∷ f} {just i1 ∷ g} {just i0 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h -3⊆∩f {nothing ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h -3⊆∩f {nothing ∷ f} {just i0 ∷ g} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h -3⊆∩f {nothing ∷ f} {just i1 ∷ g} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h -3⊆∩f {nothing ∷ f} {nothing ∷ g} {[]} f<g f<h = 3⊆∩f {f} {g} {[]} f<g f<h -3⊆∩f {nothing ∷ f} {nothing ∷ g} {just _ ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h -3⊆∩f {nothing ∷ f} {nothing ∷ g} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h - -3↑22 : (f : Nat → Two) (i j : Nat) → List (Maybe Two) -3↑22 f Zero j = [] -3↑22 f (Suc i) j = just (f j) ∷ 3↑22 f i (Suc j) - -_3↑_ : (Nat → Two) → Nat → List (Maybe Two) -_3↑_ f i = 3↑22 f i 0 - -3↑< : {f : Nat → Two} → { x y : Nat } → x ≤ y → (_3↑_ f x) 3⊆ (_3↑_ f y) -3↑< {f} {x} {y} x<y = lemma x y 0 x<y where - lemma : (x y i : Nat) → x ≤ y → (3↑22 f x i ) 3⊆ (3↑22 f y i ) - lemma 0 y i z≤n with f i - lemma Zero Zero i z≤n | i0 = refl - lemma Zero (Suc y) i z≤n | i0 = 3⊆-[] {3↑22 f (Suc y) i} - lemma Zero Zero i z≤n | i1 = refl - lemma Zero (Suc y) i z≤n | i1 = 3⊆-[] {3↑22 f (Suc y) i} - lemma (Suc x) (Suc y) i (s≤s x<y) with f i - lemma (Suc x) (Suc y) i (s≤s x<y) | i0 = lemma x y (Suc i) x<y - lemma (Suc x) (Suc y) i (s≤s x<y) | i1 = lemma x y (Suc i) x<y - -Finite3b : (p : List (Maybe Two) ) → Bool -Finite3b [] = true -Finite3b (just _ ∷ f) = Finite3b f -Finite3b (nothing ∷ f) = false - -finite3cov : (p : List (Maybe Two) ) → List (Maybe Two) -finite3cov [] = [] -finite3cov (just y ∷ x) = just y ∷ finite3cov x -finite3cov (nothing ∷ x) = just i0 ∷ finite3cov x - -Dense-3 : F-Dense (List (Maybe Two) ) (λ x → One) _3⊆_ _3∩_ -Dense-3 = record { - dense = λ x → Finite3b x ≡ true - ; d⊆P = OneObj - ; dense-f = λ x → finite3cov x - ; dense-d = λ {p} d → lemma1 p - ; dense-p = λ {p} d → lemma2 p - } where - lemma1 : (p : List (Maybe Two) ) → Finite3b (finite3cov p) ≡ true - lemma1 [] = refl - lemma1 (just i0 ∷ p) = lemma1 p - lemma1 (just i1 ∷ p) = lemma1 p - lemma1 (nothing ∷ p) = lemma1 p - lemma2 : (p : List (Maybe Two)) → p 3⊆ finite3cov p - lemma2 [] = refl - lemma2 (just i0 ∷ p) = lemma2 p - lemma2 (just i1 ∷ p) = lemma2 p - lemma2 (nothing ∷ p) = lemma2 p - --- min = Data.Nat._⊓_ --- m≤m⊔n = Data.Nat._⊔_ --- open import Data.Nat.Properties -
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/BAlgbra.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,122 @@ +open import Level +open import Ordinals +module BAlgbra {n : Level } (O : Ordinals {n}) where + +open import zf +open import logic +import OrdUtil +import OD +import ODUtil +import ODC + +open import Relation.Nullary +open import Relation.Binary +open import Data.Empty +open import Relation.Binary +open import Relation.Binary.Core +open import Relation.Binary.PropositionalEquality +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ; _+_ to _n+_ ) + +open inOrdinal O +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +open OrdUtil O +open ODUtil O + +open OD O +open OD.OD +open ODAxiom odAxiom +open HOD + +open _∧_ +open _∨_ +open Bool + +--_∩_ : ( A B : HOD ) → HOD +--A ∩ B = record { od = record { def = λ x → odef A x ∧ odef B x } ; +-- odmax = omin (odmax A) (odmax B) ; <odmax = λ y → min1 (<odmax A (proj1 y)) (<odmax B (proj2 y)) } + +_∪_ : ( A B : HOD ) → HOD +A ∪ B = record { od = record { def = λ x → odef A x ∨ odef B x } ; + odmax = omax (odmax A) (odmax B) ; <odmax = lemma } where + lemma : {y : Ordinal} → odef A y ∨ odef B y → y o< omax (odmax A) (odmax B) + lemma {y} (case1 a) = ordtrans (<odmax A a) (omax-x _ _) + lemma {y} (case2 b) = ordtrans (<odmax B b) (omax-y _ _) + +_\_ : ( A B : HOD ) → HOD +A \ B = record { od = record { def = λ x → odef A x ∧ ( ¬ ( odef B x ) ) }; odmax = odmax A ; <odmax = λ y → <odmax A (proj1 y) } + +∪-Union : { A B : HOD } → Union (A , B) ≡ ( A ∪ B ) +∪-Union {A} {B} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where + lemma1 : {x : Ordinal} → odef (Union (A , B)) x → odef (A ∪ B) x + lemma1 {x} lt = lemma3 lt where + lemma4 : {y : Ordinal} → odef (A , B) y ∧ odef (* y) x → ¬ (¬ ( odef A x ∨ odef B x) ) + lemma4 {y} z with proj1 z + lemma4 {y} z | case1 refl = double-neg (case1 ( subst (λ k → odef k x ) *iso (proj2 z)) ) + lemma4 {y} z | case2 refl = double-neg (case2 ( subst (λ k → odef k x ) *iso (proj2 z)) ) + lemma3 : (((u : Ordinal ) → ¬ odef (A , B) u ∧ odef (* u) x) → ⊥) → odef (A ∪ B) x + lemma3 not = ODC.double-neg-eilm O (FExists _ lemma4 not) -- choice + lemma2 : {x : Ordinal} → odef (A ∪ B) x → odef (Union (A , B)) x + lemma2 {x} (case1 A∋x) = subst (λ k → odef (Union (A , B)) k) &iso ( IsZF.union→ isZF (A , B) (* x) A + ⟪ case1 refl , d→∋ A A∋x ⟫ ) + lemma2 {x} (case2 B∋x) = subst (λ k → odef (Union (A , B)) k) &iso ( IsZF.union→ isZF (A , B) (* x) B + ⟪ case2 refl , d→∋ B B∋x ⟫ ) + +∩-Select : { A B : HOD } → Select A ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) ≡ ( A ∩ B ) +∩-Select {A} {B} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where + lemma1 : {x : Ordinal} → odef (Select A (λ x₁ → (A ∋ x₁) ∧ (B ∋ x₁))) x → odef (A ∩ B) x + lemma1 {x} lt = ⟪ proj1 lt , subst (λ k → odef B k ) &iso (proj2 (proj2 lt)) ⟫ + lemma2 : {x : Ordinal} → odef (A ∩ B) x → odef (Select A (λ x₁ → (A ∋ x₁) ∧ (B ∋ x₁))) x + lemma2 {x} lt = ⟪ proj1 lt , ⟪ d→∋ A (proj1 lt) , d→∋ B (proj2 lt) ⟫ ⟫ + +dist-ord : {p q r : HOD } → p ∩ ( q ∪ r ) ≡ ( p ∩ q ) ∪ ( p ∩ r ) +dist-ord {p} {q} {r} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where + lemma1 : {x : Ordinal} → odef (p ∩ (q ∪ r)) x → odef ((p ∩ q) ∪ (p ∩ r)) x + lemma1 {x} lt with proj2 lt + lemma1 {x} lt | case1 q∋x = case1 ⟪ proj1 lt , q∋x ⟫ + lemma1 {x} lt | case2 r∋x = case2 ⟪ proj1 lt , r∋x ⟫ + lemma2 : {x : Ordinal} → odef ((p ∩ q) ∪ (p ∩ r)) x → odef (p ∩ (q ∪ r)) x + lemma2 {x} (case1 p∩q) = ⟪ proj1 p∩q , case1 (proj2 p∩q ) ⟫ + lemma2 {x} (case2 p∩r) = ⟪ proj1 p∩r , case2 (proj2 p∩r ) ⟫ + +dist-ord2 : {p q r : HOD } → p ∪ ( q ∩ r ) ≡ ( p ∪ q ) ∩ ( p ∪ r ) +dist-ord2 {p} {q} {r} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where + lemma1 : {x : Ordinal} → odef (p ∪ (q ∩ r)) x → odef ((p ∪ q) ∩ (p ∪ r)) x + lemma1 {x} (case1 cp) = ⟪ case1 cp , case1 cp ⟫ + lemma1 {x} (case2 cqr) = ⟪ case2 (proj1 cqr) , case2 (proj2 cqr) ⟫ + lemma2 : {x : Ordinal} → odef ((p ∪ q) ∩ (p ∪ r)) x → odef (p ∪ (q ∩ r)) x + lemma2 {x} lt with proj1 lt | proj2 lt + lemma2 {x} lt | case1 cp | _ = case1 cp + lemma2 {x} lt | _ | case1 cp = case1 cp + lemma2 {x} lt | case2 cq | case2 cr = case2 ⟪ cq , cr ⟫ + +record IsBooleanAlgebra ( L : Set n) + ( b1 : L ) + ( b0 : L ) + ( -_ : L → L ) + ( _+_ : L → L → L ) + ( _x_ : L → L → L ) : Set (suc n) where + field + +-assoc : {a b c : L } → a + ( b + c ) ≡ (a + b) + c + x-assoc : {a b c : L } → a x ( b x c ) ≡ (a x b) x c + +-sym : {a b : L } → a + b ≡ b + a + -sym : {a b : L } → a x b ≡ b x a + +-aab : {a b : L } → a + ( a x b ) ≡ a + x-aab : {a b : L } → a x ( a + b ) ≡ a + +-dist : {a b c : L } → a + ( b x c ) ≡ ( a x b ) + ( a x c ) + x-dist : {a b c : L } → a x ( b + c ) ≡ ( a + b ) x ( a + c ) + a+0 : {a : L } → a + b0 ≡ a + ax1 : {a : L } → a x b1 ≡ a + a+-a1 : {a : L } → a + ( - a ) ≡ b1 + ax-a0 : {a : L } → a x ( - a ) ≡ b0 + +record BooleanAlgebra ( L : Set n) : Set (suc n) where + field + b1 : L + b0 : L + -_ : L → L + _+_ : L → L → L + _x_ : L → L → L + isBooleanAlgebra : IsBooleanAlgebra L b1 b0 -_ _+_ _x_ +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/LEMC.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,178 @@ +open import Level +open import Ordinals +open import logic +open import Relation.Nullary +module LEMC {n : Level } (O : Ordinals {n} ) (p∨¬p : ( p : Set n) → p ∨ ( ¬ p )) where + +open import zf +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +open import Relation.Binary.PropositionalEquality +open import Data.Nat.Properties +open import Data.Empty +open import Relation.Binary +open import Relation.Binary.Core + +open import nat +import OD + +open inOrdinal O +open OD O +open OD.OD +open OD._==_ +open ODAxiom odAxiom +import OrdUtil +import ODUtil +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +open OrdUtil O +open ODUtil O + +open import zfc + +open HOD + +open _⊆_ + +decp : ( p : Set n ) → Dec p -- assuming axiom of choice +decp p with p∨¬p p +decp p | case1 x = yes x +decp p | case2 x = no x + +∋-p : (A x : HOD ) → Dec ( A ∋ x ) +∋-p A x with p∨¬p ( A ∋ x) -- LEM +∋-p A x | case1 t = yes t +∋-p A x | case2 t = no (λ x → t x) + +double-neg-eilm : {A : Set n} → ¬ ¬ A → A -- we don't have this in intutionistic logic +double-neg-eilm {A} notnot with decp A -- assuming axiom of choice +... | yes p = p +... | no ¬p = ⊥-elim ( notnot ¬p ) + +power→⊆ : ( A t : HOD) → Power A ∋ t → t ⊆ A +power→⊆ A t PA∋t = record { incl = λ {x} t∋x → double-neg-eilm (λ not → t1 t∋x (λ x → not x) ) } where + t1 : {x : HOD } → t ∋ x → ¬ ¬ (A ∋ x) + t1 = power→ A t PA∋t + +--- With assuption of HOD is ordered, p ∨ ( ¬ p ) <=> axiom of choice +--- +record choiced ( X : Ordinal ) : Set n where + field + a-choice : Ordinal + is-in : odef (* X) a-choice + +open choiced + +-- ∋→d : ( a : HOD ) { x : HOD } → * (& a) ∋ x → X ∋ * (a-choice (choice-func X not)) +-- ∋→d a lt = subst₂ (λ j k → odef j k) *iso (sym &iso) lt + +oo∋ : { a : HOD} { x : Ordinal } → odef (* (& a)) x → a ∋ * x +oo∋ lt = subst₂ (λ j k → odef j k) *iso (sym &iso) lt + +∋oo : { a : HOD} { x : Ordinal } → a ∋ * x → odef (* (& a)) x +∋oo lt = subst₂ (λ j k → odef j k ) (sym *iso) &iso lt + +OD→ZFC : ZFC +OD→ZFC = record { + ZFSet = HOD + ; _∋_ = _∋_ + ; _≈_ = _=h=_ + ; ∅ = od∅ + ; Select = Select + ; isZFC = isZFC + } where + -- infixr 200 _∈_ + -- infixr 230 _∩_ _∪_ + isZFC : IsZFC (HOD ) _∋_ _=h=_ od∅ Select + isZFC = record { + choice-func = λ A {X} not A∋X → * (a-choice (choice-func X not) ); + choice = λ A {X} A∋X not → oo∋ (is-in (choice-func X not)) + } where + -- + -- the axiom choice from LEM and OD ordering + -- + choice-func : (X : HOD ) → ¬ ( X =h= od∅ ) → choiced (& X) + choice-func X not = have_to_find where + ψ : ( ox : Ordinal ) → Set n + ψ ox = (( x : Ordinal ) → x o< ox → ( ¬ odef X x )) ∨ choiced (& X) + lemma-ord : ( ox : Ordinal ) → ψ ox + lemma-ord ox = TransFinite0 {ψ} induction ox where + ∀-imply-or : {A : Ordinal → Set n } {B : Set n } + → ((x : Ordinal ) → A x ∨ B) → ((x : Ordinal ) → A x) ∨ B + ∀-imply-or {A} {B} ∀AB with p∨¬p ((x : Ordinal ) → A x) -- LEM + ∀-imply-or {A} {B} ∀AB | case1 t = case1 t + ∀-imply-or {A} {B} ∀AB | case2 x = case2 (lemma (λ not → x not )) where + lemma : ¬ ((x : Ordinal ) → A x) → B + lemma not with p∨¬p B + lemma not | case1 b = b + lemma not | case2 ¬b = ⊥-elim (not (λ x → dont-orb (∀AB x) ¬b )) + induction : (x : Ordinal) → ((y : Ordinal) → y o< x → ψ y) → ψ x + induction x prev with ∋-p X ( * x) + ... | yes p = case2 ( record { a-choice = x ; is-in = ∋oo p } ) + ... | no ¬p = lemma where + lemma1 : (y : Ordinal) → (y o< x → odef X y → ⊥) ∨ choiced (& X) + lemma1 y with ∋-p X (* y) + lemma1 y | yes y<X = case2 ( record { a-choice = y ; is-in = ∋oo y<X } ) + lemma1 y | no ¬y<X = case1 ( λ lt y<X → ¬y<X (d→∋ X y<X) ) + lemma : ((y : Ordinal) → y o< x → odef X y → ⊥) ∨ choiced (& X) + lemma = ∀-imply-or lemma1 + odef→o< : {X : HOD } → {x : Ordinal } → odef X x → x o< & X + odef→o< {X} {x} lt = o<-subst {_} {_} {x} {& X} ( c<→o< ( subst₂ (λ j k → odef j k ) (sym *iso) (sym &iso) lt )) &iso &iso + have_to_find : choiced (& X) + have_to_find = dont-or (lemma-ord (& X )) ¬¬X∋x where + ¬¬X∋x : ¬ ((x : Ordinal) → x o< (& X) → odef X x → ⊥) + ¬¬X∋x nn = not record { + eq→ = λ {x} lt → ⊥-elim (nn x (odef→o< lt) lt) + ; eq← = λ {x} lt → ⊥-elim ( ¬x<0 lt ) + } + + -- + -- axiom regurality from ε-induction (using axiom of choice above) + -- + -- from https://math.stackexchange.com/questions/2973777/is-it-possible-to-prove-regularity-with-transfinite-induction-only + -- + -- FIXME : don't use HOD make this level n, so we can remove ε-induction1 + record Minimal (x : HOD) : Set (suc n) where + field + min : HOD + x∋min : x ∋ min + min-empty : (y : HOD ) → ¬ ( min ∋ y) ∧ (x ∋ y) + open Minimal + open _∧_ + induction : {x : HOD} → ({y : HOD} → x ∋ y → (u : HOD ) → (u∋x : u ∋ y) → Minimal u ) + → (u : HOD ) → (u∋x : u ∋ x) → Minimal u + induction {x} prev u u∋x with p∨¬p ((y : Ordinal ) → ¬ (odef x y) ∧ (odef u y)) + ... | case1 P = + record { min = x + ; x∋min = u∋x + ; min-empty = λ y → P (& y) + } + ... | case2 NP = min2 where + p : HOD + p = record { od = record { def = λ y1 → odef x y1 ∧ odef u y1 } ; odmax = omin (odmax x) (odmax u) ; <odmax = lemma } where + lemma : {y : Ordinal} → OD.def (od x) y ∧ OD.def (od u) y → y o< omin (odmax x) (odmax u) + lemma {y} lt = min1 (<odmax x (proj1 lt)) (<odmax u (proj2 lt)) + np : ¬ (p =h= od∅) + np p∅ = NP (λ y p∋y → ∅< {p} {_} (d→∋ p p∋y) p∅ ) + y1choice : choiced (& p) + y1choice = choice-func p np + y1 : HOD + y1 = * (a-choice y1choice) + y1prop : (x ∋ y1) ∧ (u ∋ y1) + y1prop = oo∋ (is-in y1choice) + min2 : Minimal u + min2 = prev (proj1 y1prop) u (proj2 y1prop) + Min2 : (x : HOD) → (u : HOD ) → (u∋x : u ∋ x) → Minimal u + Min2 x u u∋x = (ε-induction {λ y → (u : HOD ) → (u∋x : u ∋ y) → Minimal u } induction x u u∋x ) + cx : {x : HOD} → ¬ (x =h= od∅ ) → choiced (& x ) + cx {x} nx = choice-func x nx + minimal : (x : HOD ) → ¬ (x =h= od∅ ) → HOD + minimal x ne = min (Min2 (* (a-choice (cx {x} ne) )) x ( oo∋ (is-in (cx ne))) ) + x∋minimal : (x : HOD ) → ( ne : ¬ (x =h= od∅ ) ) → odef x ( & ( minimal x ne ) ) + x∋minimal x ne = x∋min (Min2 (* (a-choice (cx {x} ne) )) x ( oo∋ (is-in (cx ne))) ) + minimal-1 : (x : HOD ) → ( ne : ¬ (x =h= od∅ ) ) → (y : HOD ) → ¬ ( odef (minimal x ne) (& y)) ∧ (odef x (& y) ) + minimal-1 x ne y = min-empty (Min2 (* (a-choice (cx ne) )) x ( oo∋ (is-in (cx ne)))) y + + + +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/OD.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,618 @@ +{-# OPTIONS --allow-unsolved-metas #-} +open import Level +open import Ordinals +module OD {n : Level } (O : Ordinals {n} ) where + +open import zf +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +open import Relation.Binary.PropositionalEquality hiding ( [_] ) +open import Data.Nat.Properties +open import Data.Empty +open import Relation.Nullary +open import Relation.Binary hiding (_⇔_) +open import Relation.Binary.Core hiding (_⇔_) + +open import logic +import OrdUtil +open import nat + +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +open OrdUtil O + +-- Ordinal Definable Set + +record OD : Set (suc n ) where + field + def : (x : Ordinal ) → Set n + +open 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 _==_ + +==-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 + +-- next assumptions are our axiom +-- +-- 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. +-- +-- ==→o≡ is necessary to prove axiom of extensionality. + +-- Ordinals in OD , the maximum +Ords : OD +Ords = record { def = λ x → One } + +record HOD : Set (suc n) where + field + od : OD + odmax : Ordinal + <odmax : {y : Ordinal} → def od y → y o< odmax + +open HOD + +record ODAxiom : Set (suc n) where + field + -- HOD is isomorphic to Ordinal (by means of Goedel number) + & : HOD → Ordinal + * : Ordinal → HOD + c<→o< : {x y : HOD } → def (od y) ( & x ) → & x o< & y + ⊆→o≤ : {y z : HOD } → ({x : Ordinal} → def (od y) x → def (od z) x ) → & y o< osuc (& z) + *iso : {x : HOD } → * ( & x ) ≡ x + &iso : {x : Ordinal } → & ( * x ) ≡ x + ==→o≡ : {x y : HOD } → (od x == od y) → x ≡ y + sup-o : (A : HOD) → ( ( x : Ordinal ) → def (od A) x → Ordinal ) → Ordinal + sup-o< : (A : HOD) → { ψ : ( x : Ordinal ) → def (od A) x → Ordinal } → ∀ {x : Ordinal } → (lt : def (od A) x ) → ψ x lt o< sup-o A ψ +-- possible order restriction + ho< : {x : HOD} → & x o< next (odmax x) + + +postulate odAxiom : ODAxiom +open ODAxiom odAxiom + +-- odmax minimality +-- +-- since we have ==→o≡ , so odmax have to be unique. We should have odmaxmin in HOD. +-- We can calculate the minimum using sup but it is tedius. +-- Only Select has non minimum odmax. +-- We have the same problem on 'def' itself, but we leave it. + +odmaxmin : Set (suc n) +odmaxmin = (y : HOD) (z : Ordinal) → ((x : Ordinal)→ def (od y) x → x o< z) → odmax y o< osuc z + +-- OD ⇔ Ordinal leads a contradiction, so we need bounded HOD +¬OD-order : ( & : OD → Ordinal ) → ( * : Ordinal → OD ) → ( { x y : OD } → def y ( & x ) → & x o< & y) → ⊥ +¬OD-order & * c<→o< = osuc-< <-osuc (c<→o< {Ords} OneObj ) + +-- Open supreme upper bound leads a contradition, so we use domain restriction on sup +¬open-sup : ( sup-o : (Ordinal → Ordinal ) → Ordinal) → ((ψ : Ordinal → Ordinal ) → (x : Ordinal) → ψ x o< sup-o ψ ) → ⊥ +¬open-sup sup-o sup-o< = o<> <-osuc (sup-o< next-ord (sup-o next-ord)) where + next-ord : Ordinal → Ordinal + next-ord x = osuc x + +-- Ordinal in OD ( and ZFSet ) Transitive Set +Ord : ( a : Ordinal ) → HOD +Ord a = record { od = record { def = λ y → y o< a } ; odmax = a ; <odmax = lemma } where + lemma : {x : Ordinal} → x o< a → x o< a + lemma {x} lt = lt + +od∅ : HOD +od∅ = Ord o∅ + +odef : HOD → Ordinal → Set n +odef A x = def ( od A ) x + +_∋_ : ( a x : HOD ) → Set n +_∋_ a x = odef a ( & x ) + +-- _c<_ : ( x a : HOD ) → Set n +-- x c< a = a ∋ x + +d→∋ : ( a : HOD ) { x : Ordinal} → odef a x → a ∋ (* x) +d→∋ a lt = subst (λ k → odef a k ) (sym &iso) lt + +-- odef-subst : {Z : HOD } {X : Ordinal }{z : HOD } {x : Ordinal }→ odef Z X → Z ≡ z → X ≡ x → odef z x +-- odef-subst df refl refl = df + +otrans : {a x y : Ordinal } → odef (Ord a) x → odef (Ord x) y → odef (Ord a) y +otrans x<a y<x = ordtrans y<x x<a + +-- If we have reverse of c<→o<, everything becomes Ordinal +∈→c<→HOD=Ord : ( o<→c< : {x y : Ordinal } → x o< y → odef (* y) x ) → {x : HOD } → x ≡ Ord (& x) +∈→c<→HOD=Ord o<→c< {x} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where + lemma1 : {y : Ordinal} → odef x y → odef (Ord (& x)) y + lemma1 {y} lt = subst ( λ k → k o< & x ) &iso (c<→o< {* y} {x} (d→∋ x lt)) + lemma2 : {y : Ordinal} → odef (Ord (& x)) y → odef x y + lemma2 {y} lt = subst (λ k → odef k y ) *iso (o<→c< {y} {& x} lt ) + +-- avoiding lv != Zero error +orefl : { x : HOD } → { y : Ordinal } → & x ≡ y → & x ≡ y +orefl refl = refl + +==-iso : { x y : HOD } → od (* (& x)) == od (* (& y)) → od x == od y +==-iso {x} {y} eq = record { + eq→ = λ {z} d → lemma ( eq→ eq (subst (λ k → odef k z ) (sym *iso) d )) ; + eq← = λ {z} d → lemma ( eq← eq (subst (λ k → odef k z ) (sym *iso) d )) } + where + lemma : {x : HOD } {z : Ordinal } → odef (* (& x)) z → odef x z + lemma {x} {z} d = subst (λ k → odef k z) (*iso) d + +=-iso : {x y : HOD } → (od x == od y) ≡ (od (* (& x)) == od y) +=-iso {_} {y} = cong ( λ k → od k == od y ) (sym *iso) + +ord→== : { x y : HOD } → & x ≡ & y → od x == od y +ord→== {x} {y} eq = ==-iso (lemma (& x) (& y) (orefl eq)) where + lemma : ( ox oy : Ordinal ) → ox ≡ oy → od (* ox) == od (* oy) + lemma ox ox refl = ==-refl + +o≡→== : { x y : Ordinal } → x ≡ y → od (* x) == od (* y) +o≡→== {x} {.x} refl = ==-refl + +o∅≡od∅ : * (o∅ ) ≡ od∅ +o∅≡od∅ = ==→o≡ lemma where + lemma0 : {x : Ordinal} → odef (* o∅) x → odef od∅ x + lemma0 {x} lt with c<→o< {* x} {* o∅} (subst (λ k → odef (* o∅) k ) (sym &iso) lt) + ... | t = subst₂ (λ j k → j o< k ) &iso &iso t + lemma1 : {x : Ordinal} → odef od∅ x → odef (* o∅) x + lemma1 {x} lt = ⊥-elim (¬x<0 lt) + lemma : od (* o∅) == od od∅ + lemma = record { eq→ = lemma0 ; eq← = lemma1 } + +ord-od∅ : & (od∅ ) ≡ o∅ +ord-od∅ = sym ( subst (λ k → k ≡ & (od∅ ) ) &iso (cong ( λ k → & k ) o∅≡od∅ ) ) + +≡o∅→=od∅ : {x : HOD} → & x ≡ o∅ → od x == od od∅ +≡o∅→=od∅ {x} eq = record { eq→ = λ {y} lt → ⊥-elim ( ¬x<0 {y} (subst₂ (λ j k → j o< k ) &iso eq ( c<→o< {* y} {x} (d→∋ x lt)))) + ; eq← = λ {y} lt → ⊥-elim ( ¬x<0 lt )} + +=od∅→≡o∅ : {x : HOD} → od x == od od∅ → & x ≡ o∅ +=od∅→≡o∅ {x} eq = trans (cong (λ k → & k ) (==→o≡ {x} {od∅} eq)) ord-od∅ + +∅0 : record { def = λ x → Lift n ⊥ } == od od∅ +eq→ ∅0 {w} (lift ()) +eq← ∅0 {w} lt = lift (¬x<0 lt) + +∅< : { x y : HOD } → odef x (& y ) → ¬ ( od x == od od∅ ) +∅< {x} {y} d eq with eq→ (==-trans eq (==-sym ∅0) ) d +∅< {x} {y} d eq | lift () + +∅6 : { x : HOD } → ¬ ( x ∋ x ) -- no Russel paradox +∅6 {x} x∋x = o<¬≡ refl ( c<→o< {x} {x} x∋x ) + +odef-iso : {A B : HOD } {x y : Ordinal } → x ≡ y → (odef A y → odef B y) → odef A x → odef B x +odef-iso refl t = t + +is-o∅ : ( x : Ordinal ) → Dec ( x ≡ o∅ ) +is-o∅ x with trio< x o∅ +is-o∅ x | tri< a ¬b ¬c = no ¬b +is-o∅ x | tri≈ ¬a b ¬c = yes b +is-o∅ x | tri> ¬a ¬b c = no ¬b + +-- the pair +_,_ : HOD → HOD → HOD +x , y = record { od = record { def = λ t → (t ≡ & x ) ∨ ( t ≡ & y ) } ; odmax = omax (& x) (& y) ; <odmax = lemma } where + lemma : {t : Ordinal} → (t ≡ & x) ∨ (t ≡ & y) → t o< omax (& x) (& y) + lemma {t} (case1 refl) = omax-x _ _ + lemma {t} (case2 refl) = omax-y _ _ + +pair<y : {x y : HOD } → y ∋ x → & (x , x) o< osuc (& y) +pair<y {x} {y} y∋x = ⊆→o≤ lemma where + lemma : {z : Ordinal} → def (od (x , x)) z → def (od y) z + lemma (case1 refl) = y∋x + lemma (case2 refl) = y∋x + +-- another possible restriction. We reqest no minimality on odmax, so it may arbitrary larger. +odmax<& : { x y : HOD } → x ∋ y → Set n +odmax<& {x} {y} x∋y = odmax x o< & x + +in-codomain : (X : HOD ) → ( ψ : HOD → HOD ) → OD +in-codomain X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( odef X y ∧ ( x ≡ & (ψ (* y ))))) } + +_∩_ : ( A B : HOD ) → HOD +A ∩ B = record { od = record { def = λ x → odef A x ∧ odef B x } + ; odmax = omin (odmax A) (odmax B) ; <odmax = λ y → min1 (<odmax A (proj1 y)) (<odmax B (proj2 y))} + +record _⊆_ ( A B : HOD ) : Set (suc n) where + field + incl : { x : HOD } → A ∋ x → B ∋ x + +open _⊆_ +infixr 220 _⊆_ + +-- if we have & (x , x) ≡ osuc (& x), ⊆→o≤ → c<→o< +⊆→o≤→c<→o< : ({x : HOD} → & (x , x) ≡ osuc (& x) ) + → ({y z : HOD } → ({x : Ordinal} → def (od y) x → def (od z) x ) → & y o< osuc (& z) ) + → {x y : HOD } → def (od y) ( & x ) → & x o< & y +⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x with trio< (& x) (& y) +⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri< a ¬b ¬c = a +⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri≈ ¬a b ¬c = ⊥-elim ( o<¬≡ (peq {x}) (pair<y (subst (λ k → k ∋ x) (sym ( ==→o≡ {x} {y} (ord→== b))) y∋x ))) +⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri> ¬a ¬b c = + ⊥-elim ( o<> (⊆→o≤ {x , x} {y} y⊆x,x ) lemma1 ) where + lemma : {z : Ordinal} → (z ≡ & x) ∨ (z ≡ & x) → & x ≡ z + lemma (case1 refl) = refl + lemma (case2 refl) = refl + y⊆x,x : {z : Ordinal} → def (od (x , x)) z → def (od y) z + y⊆x,x {z} lt = subst (λ k → def (od y) k ) (lemma lt) y∋x + lemma1 : osuc (& y) o< & (x , x) + lemma1 = subst (λ k → osuc (& y) o< k ) (sym (peq {x})) (osucc c ) + +ε-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 + +Select : (X : HOD ) → ((x : HOD ) → Set n ) → HOD +Select X ψ = record { od = record { def = λ x → ( odef X x ∧ ψ ( * x )) } ; odmax = odmax X ; <odmax = λ y → <odmax X (proj1 y) } + +Replace : HOD → (HOD → HOD) → HOD +Replace X ψ = record { od = record { def = λ x → (x o< sup-o X (λ y X∋y → & (ψ (* y)))) ∧ def (in-codomain X ψ) x } + ; odmax = rmax ; <odmax = rmax<} where + rmax : Ordinal + rmax = sup-o X (λ y X∋y → & (ψ (* y))) + rmax< : {y : Ordinal} → (y o< rmax) ∧ def (in-codomain X ψ) y → y o< rmax + rmax< lt = proj1 lt + +-- +-- If we have LEM, Replace' is equivalent to Replace +-- +in-codomain' : (X : HOD ) → ((x : HOD) → X ∋ x → HOD) → OD +in-codomain' X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( odef X y ∧ ((lt : odef X y) → x ≡ & (ψ (* y ) (d→∋ X lt) )))) } +Replace' : (X : HOD) → ((x : HOD) → X ∋ x → HOD) → HOD +Replace' X ψ = record { od = record { def = λ x → (x o< sup-o X (λ y X∋y → & (ψ (* y) (d→∋ X X∋y) ))) ∧ def (in-codomain' X ψ) x } + ; odmax = rmax ; <odmax = rmax< } where + rmax : Ordinal + rmax = sup-o X (λ y X∋y → & (ψ (* y) (d→∋ X X∋y))) + rmax< : {y : Ordinal} → (y o< rmax) ∧ def (in-codomain' X ψ) y → y o< rmax + rmax< lt = proj1 lt + +Union : HOD → HOD +Union U = record { od = record { def = λ x → ¬ (∀ (u : Ordinal ) → ¬ ((odef U u) ∧ (odef (* u) x))) } + ; odmax = osuc (& U) ; <odmax = umax< } where + umax< : {y : Ordinal} → ¬ ((u : Ordinal) → ¬ def (od U) u ∧ def (od (* u)) y) → y o< osuc (& U) + umax< {y} not = lemma (FExists _ lemma1 not ) where + lemma0 : {x : Ordinal} → def (od (* x)) y → y o< x + lemma0 {x} x<y = subst₂ (λ j k → j o< k ) &iso &iso (c<→o< (d→∋ (* x) x<y )) + lemma2 : {x : Ordinal} → def (od U) x → x o< & U + lemma2 {x} x<U = subst (λ k → k o< & U ) &iso (c<→o< (d→∋ U x<U)) + lemma1 : {x : Ordinal} → def (od U) x ∧ def (od (* x)) y → ¬ (& U o< y) + lemma1 {x} lt u<y = o<> u<y (ordtrans (lemma0 (proj2 lt)) (lemma2 (proj1 lt)) ) + lemma : ¬ ((& U) o< y ) → y o< osuc (& U) + lemma not with trio< y (& U) + lemma not | tri< a ¬b ¬c = ordtrans a <-osuc + lemma not | tri≈ ¬a refl ¬c = <-osuc + lemma not | tri> ¬a ¬b c = ⊥-elim (not c) +_∈_ : ( A B : HOD ) → Set n +A ∈ B = B ∋ A + +OPwr : (A : HOD ) → HOD +OPwr A = Ord ( sup-o (Ord (osuc (& A))) ( λ x A∋x → & ( A ∩ (* x)) ) ) + +Power : HOD → HOD +Power A = Replace (OPwr (Ord (& A))) ( λ x → A ∩ x ) +-- {_} : ZFSet → ZFSet +-- { x } = ( x , x ) -- better to use (x , x) directly + +union→ : (X z u : HOD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z +union→ X z u xx not = ⊥-elim ( not (& u) ( ⟪ proj1 xx + , subst ( λ k → odef k (& z)) (sym *iso) (proj2 xx) ⟫ )) +union← : (X z : HOD) (X∋z : Union X ∋ z) → ¬ ( (u : HOD ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) +union← X z UX∋z = FExists _ lemma UX∋z where + lemma : {y : Ordinal} → odef X y ∧ odef (* y) (& z) → ¬ ((u : HOD) → ¬ (X ∋ u) ∧ (u ∋ z)) + lemma {y} xx not = not (* y) ⟪ d→∋ X (proj1 xx) , proj2 xx ⟫ + +data infinite-d : ( x : Ordinal ) → Set n where + iφ : infinite-d o∅ + isuc : {x : Ordinal } → infinite-d x → + infinite-d (& ( Union (* x , (* x , * x ) ) )) + +-- ω can be diverged in our case, since we have no restriction on the corresponding ordinal of a pair. +-- We simply assumes infinite-d y has a maximum. +-- +-- This means that many of OD may not be HODs because of the & mapping divergence. +-- We should have some axioms to prevent this such as & x o< next (odmax x). +-- +-- postulate +-- ωmax : Ordinal +-- <ωmax : {y : Ordinal} → infinite-d y → y o< ωmax +-- +-- infinite : HOD +-- infinite = record { od = record { def = λ x → infinite-d x } ; odmax = ωmax ; <odmax = <ωmax } + +infinite : HOD +infinite = record { od = record { def = λ x → infinite-d x } ; odmax = next o∅ ; <odmax = lemma } where + u : (y : Ordinal ) → HOD + u y = Union (* y , (* y , * y)) + -- next< : {x y z : Ordinal} → x o< next z → y o< next x → y o< next z + lemma8 : {y : Ordinal} → & (* y , * y) o< next (odmax (* y , * y)) + lemma8 = ho< + --- (x,y) < next (omax x y) < next (osuc y) = next y + lemmaa : {x y : HOD} → & x o< & y → & (x , y) o< next (& y) + lemmaa {x} {y} x<y = subst (λ k → & (x , y) o< k ) (sym nexto≡) (subst (λ k → & (x , y) o< next k ) (sym (omax< _ _ x<y)) ho< ) + lemma81 : {y : Ordinal} → & (* y , * y) o< next (& (* y)) + lemma81 {y} = nexto=n (subst (λ k → & (* y , * y) o< k ) (cong (λ k → next k) (omxx _)) lemma8) + lemma9 : {y : Ordinal} → & (* y , (* y , * y)) o< next (& (* y , * y)) + lemma9 = lemmaa (c<→o< (case1 refl)) + lemma71 : {y : Ordinal} → & (* y , (* y , * y)) o< next (& (* y)) + lemma71 = next< lemma81 lemma9 + lemma1 : {y : Ordinal} → & (u y) o< next (osuc (& (* y , (* y , * y)))) + lemma1 = ho< + --- main recursion + lemma : {y : Ordinal} → infinite-d y → y o< next o∅ + lemma {o∅} iφ = x<nx + lemma (isuc {y} x) = next< (lemma x) (next< (subst (λ k → & (* y , (* y , * y)) o< next k) &iso lemma71 ) (nexto=n lemma1)) + +empty : (x : HOD ) → ¬ (od∅ ∋ x) +empty x = ¬x<0 + +_=h=_ : (x y : HOD) → Set n +x =h= y = od x == od y + +pair→ : ( x y t : HOD ) → (x , y) ∋ t → ( t =h= x ) ∨ ( t =h= y ) +pair→ x y t (case1 t≡x ) = case1 (subst₂ (λ j k → j =h= k ) *iso *iso (o≡→== t≡x )) +pair→ x y t (case2 t≡y ) = case2 (subst₂ (λ j k → j =h= k ) *iso *iso (o≡→== t≡y )) + +pair← : ( x y t : HOD ) → ( t =h= x ) ∨ ( t =h= y ) → (x , y) ∋ t +pair← x y t (case1 t=h=x) = case1 (cong (λ k → & k ) (==→o≡ t=h=x)) +pair← x y t (case2 t=h=y) = case2 (cong (λ k → & k ) (==→o≡ t=h=y)) + +o<→c< : {x y : Ordinal } → x o< y → (Ord x) ⊆ (Ord y) +o<→c< lt = record { incl = λ z → ordtrans z lt } + +⊆→o< : {x y : Ordinal } → (Ord x) ⊆ (Ord y) → x o< osuc y +⊆→o< {x} {y} lt with trio< x y +⊆→o< {x} {y} lt | tri< a ¬b ¬c = ordtrans a <-osuc +⊆→o< {x} {y} lt | tri≈ ¬a b ¬c = subst ( λ k → k o< osuc y) (sym b) <-osuc +⊆→o< {x} {y} lt | tri> ¬a ¬b c with (incl lt) (o<-subst c (sym &iso) refl ) +... | ttt = ⊥-elim ( o<¬≡ refl (o<-subst ttt &iso refl )) + +ψiso : {ψ : HOD → Set n} {x y : HOD } → ψ x → x ≡ y → ψ y +ψiso {ψ} t refl = t +selection : {ψ : HOD → Set n} {X y : HOD} → ((X ∋ y) ∧ ψ y) ⇔ (Select X ψ ∋ y) +selection {ψ} {X} {y} = ⟪ + ( λ cond → ⟪ proj1 cond , ψiso {ψ} (proj2 cond) (sym *iso) ⟫ ) + , ( λ select → ⟪ proj1 select , ψiso {ψ} (proj2 select) *iso ⟫ ) + ⟫ + +selection-in-domain : {ψ : HOD → Set n} {X y : HOD} → Select X ψ ∋ y → X ∋ y +selection-in-domain {ψ} {X} {y} lt = proj1 ((proj2 (selection {ψ} {X} )) lt) + +sup-c< : (ψ : HOD → HOD) → {X x : HOD} → X ∋ x → & (ψ x) o< (sup-o X (λ y X∋y → & (ψ (* y)))) +sup-c< ψ {X} {x} lt = subst (λ k → & (ψ k) o< _ ) *iso (sup-o< X lt ) +replacement← : {ψ : HOD → HOD} (X x : HOD) → X ∋ x → Replace X ψ ∋ ψ x +replacement← {ψ} X x lt = ⟪ sup-c< ψ {X} {x} lt , lemma ⟫ where + lemma : def (in-codomain X ψ) (& (ψ x)) + lemma not = ⊥-elim ( not ( & x ) ⟪ lt , cong (λ k → & (ψ k)) (sym *iso)⟫ ) +replacement→ : {ψ : HOD → HOD} (X x : HOD) → (lt : Replace X ψ ∋ x) → ¬ ( (y : HOD) → ¬ (x =h= ψ y)) +replacement→ {ψ} X x lt = contra-position lemma (lemma2 (proj2 lt)) where + lemma2 : ¬ ((y : Ordinal) → ¬ odef X y ∧ ((& x) ≡ & (ψ (* y)))) + → ¬ ((y : Ordinal) → ¬ odef X y ∧ (* (& x) =h= ψ (* y))) + lemma2 not not2 = not ( λ y d → not2 y ⟪ proj1 d , lemma3 (proj2 d)⟫) where + lemma3 : {y : Ordinal } → (& x ≡ & (ψ (* y))) → (* (& x) =h= ψ (* y)) + lemma3 {y} eq = subst (λ k → * (& x) =h= k ) *iso (o≡→== eq ) + lemma : ( (y : HOD) → ¬ (x =h= ψ y)) → ( (y : Ordinal) → ¬ odef X y ∧ (* (& x) =h= ψ (* y)) ) + lemma not y not2 = not (* y) (subst (λ k → k =h= ψ (* y)) *iso ( proj2 not2 )) + +--- +--- Power Set +--- +--- First consider ordinals in HOD +--- +--- A ∩ x = record { def = λ y → odef A y ∧ odef x y } subset of A +-- +-- +∩-≡ : { a b : HOD } → ({x : HOD } → (a ∋ x → b ∋ x)) → a =h= ( b ∩ a ) +∩-≡ {a} {b} inc = record { + eq→ = λ {x} x<a → ⟪ (subst (λ k → odef b k ) &iso (inc (d→∋ a x<a))) , x<a ⟫ ; + eq← = λ {x} x<a∩b → proj2 x<a∩b } +-- +-- Transitive Set case +-- we have t ∋ x → Ord a ∋ x means t is a subset of Ord a, that is (Ord a) ∩ t =h= t +-- OPwr (Ord a) is a sup of (Ord a) ∩ t, so OPwr (Ord a) ∋ t +-- OPwr A = Ord ( sup-o ( λ x → & ( A ∩ (* x )) ) ) +-- +ord-power← : (a : Ordinal ) (t : HOD) → ({x : HOD} → (t ∋ x → (Ord a) ∋ x)) → OPwr (Ord a) ∋ t +ord-power← a t t→A = subst (λ k → odef (OPwr (Ord a)) k ) (lemma1 lemma-eq) lemma where + lemma-eq : ((Ord a) ∩ t) =h= t + eq→ lemma-eq {z} w = proj2 w + eq← lemma-eq {z} w = ⟪ subst (λ k → odef (Ord a) k ) &iso ( t→A (d→∋ t w)) , w ⟫ + lemma1 : {a : Ordinal } { t : HOD } + → (eq : ((Ord a) ∩ t) =h= t) → & ((Ord a) ∩ (* (& t))) ≡ & t + lemma1 {a} {t} eq = subst (λ k → & ((Ord a) ∩ k) ≡ & t ) (sym *iso) (cong (λ k → & k ) (==→o≡ eq )) + lemma2 : (& t) o< (osuc (& (Ord a))) + lemma2 = ⊆→o≤ {t} {Ord a} (λ {x} x<t → subst (λ k → def (od (Ord a)) k) &iso (t→A (d→∋ t x<t))) + lemma : & ((Ord a) ∩ (* (& t)) ) o< sup-o (Ord (osuc (& (Ord a)))) (λ x lt → & ((Ord a) ∩ (* x))) + lemma = sup-o< _ lemma2 + +-- +-- Every set in HOD is a subset of Ordinals, so make OPwr (Ord (& A)) first +-- then replace of all elements of the Power set by A ∩ y +-- +-- Power A = Replace (OPwr (Ord (& A))) ( λ y → A ∩ y ) + +-- we have oly double negation form because of the replacement axiom +-- +power→ : ( A t : HOD) → Power A ∋ t → {x : HOD} → t ∋ x → ¬ ¬ (A ∋ x) +power→ A t P∋t {x} t∋x = FExists _ lemma5 lemma4 where + a = & A + lemma2 : ¬ ( (y : HOD) → ¬ (t =h= (A ∩ y))) + lemma2 = replacement→ {λ x → A ∩ x} (OPwr (Ord (& A))) t P∋t + lemma3 : (y : HOD) → t =h= ( A ∩ y ) → ¬ ¬ (A ∋ x) + lemma3 y eq not = not (proj1 (eq→ eq t∋x)) + lemma4 : ¬ ((y : Ordinal) → ¬ (t =h= (A ∩ * y))) + lemma4 not = lemma2 ( λ y not1 → not (& y) (subst (λ k → t =h= ( A ∩ k )) (sym *iso) not1 )) + lemma5 : {y : Ordinal} → t =h= (A ∩ * y) → ¬ ¬ (odef A (& x)) + lemma5 {y} eq not = (lemma3 (* y) eq) not + +power← : (A t : HOD) → ({x : HOD} → (t ∋ x → A ∋ x)) → Power A ∋ t +power← A t t→A = ⟪ lemma1 , lemma2 ⟫ where + a = & A + lemma0 : {x : HOD} → t ∋ x → Ord a ∋ x + lemma0 {x} t∋x = c<→o< (t→A t∋x) + lemma3 : OPwr (Ord a) ∋ t + lemma3 = ord-power← a t lemma0 + lemma4 : (A ∩ * (& t)) ≡ t + lemma4 = let open ≡-Reasoning in begin + A ∩ * (& t) + ≡⟨ cong (λ k → A ∩ k) *iso ⟩ + A ∩ t + ≡⟨ sym (==→o≡ ( ∩-≡ {t} {A} t→A )) ⟩ + t + ∎ + sup1 : Ordinal + sup1 = sup-o (Ord (osuc (& (Ord (& A))))) (λ x A∋x → & ((Ord (& A)) ∩ (* x))) + lemma9 : def (od (Ord (Ordinals.osuc O (& (Ord (& A)))))) (& (Ord (& A))) + lemma9 = <-osuc + lemmab : & ((Ord (& A)) ∩ (* (& (Ord (& A) )))) o< sup1 + lemmab = sup-o< (Ord (osuc (& (Ord (& A))))) lemma9 + lemmad : Ord (osuc (& A)) ∋ t + lemmad = ⊆→o≤ (λ {x} lt → subst (λ k → def (od A) k ) &iso (t→A (d→∋ t lt))) + lemmac : ((Ord (& A)) ∩ (* (& (Ord (& A) )))) =h= Ord (& A) + lemmac = record { eq→ = lemmaf ; eq← = lemmag } where + lemmaf : {x : Ordinal} → def (od ((Ord (& A)) ∩ (* (& (Ord (& A)))))) x → def (od (Ord (& A))) x + lemmaf {x} lt = proj1 lt + lemmag : {x : Ordinal} → def (od (Ord (& A))) x → def (od ((Ord (& A)) ∩ (* (& (Ord (& A)))))) x + lemmag {x} lt = ⟪ lt , subst (λ k → def (od k) x) (sym *iso) lt ⟫ + lemmae : & ((Ord (& A)) ∩ (* (& (Ord (& A))))) ≡ & (Ord (& A)) + lemmae = cong (λ k → & k ) ( ==→o≡ lemmac) + lemma7 : def (od (OPwr (Ord (& A)))) (& t) + lemma7 with osuc-≡< lemmad + lemma7 | case2 lt = ordtrans (c<→o< lt) (subst (λ k → k o< sup1) lemmae lemmab ) + lemma7 | case1 eq with osuc-≡< (⊆→o≤ {* (& t)} {* (& (Ord (& t)))} (λ {x} lt → lemmah lt )) where + lemmah : {x : Ordinal } → def (od (* (& t))) x → def (od (* (& (Ord (& t))))) x + lemmah {x} lt = subst (λ k → def (od k) x ) (sym *iso) (subst (λ k → k o< (& t)) + &iso + (c<→o< (subst₂ (λ j k → def (od j) k) *iso (sym &iso) lt ))) + lemma7 | case1 eq | case1 eq1 = subst (λ k → k o< sup1) (trans lemmae lemmai) lemmab where + lemmai : & (Ord (& A)) ≡ & t + lemmai = let open ≡-Reasoning in begin + & (Ord (& A)) + ≡⟨ sym (cong (λ k → & (Ord k)) eq) ⟩ + & (Ord (& t)) + ≡⟨ sym &iso ⟩ + & (* (& (Ord (& t)))) + ≡⟨ sym eq1 ⟩ + & (* (& t)) + ≡⟨ &iso ⟩ + & t + ∎ + lemma7 | case1 eq | case2 lt = ordtrans lemmaj (subst (λ k → k o< sup1) lemmae lemmab ) where + lemmak : & (* (& (Ord (& t)))) ≡ & (Ord (& A)) + lemmak = let open ≡-Reasoning in begin + & (* (& (Ord (& t)))) + ≡⟨ &iso ⟩ + & (Ord (& t)) + ≡⟨ cong (λ k → & (Ord k)) eq ⟩ + & (Ord (& A)) + ∎ + lemmaj : & t o< & (Ord (& A)) + lemmaj = subst₂ (λ j k → j o< k ) &iso lemmak lt + lemma1 : & t o< sup-o (OPwr (Ord (& A))) (λ x lt → & (A ∩ (* x))) + lemma1 = subst (λ k → & k o< sup-o (OPwr (Ord (& A))) (λ x lt → & (A ∩ (* x)))) + lemma4 (sup-o< (OPwr (Ord (& A))) lemma7 ) + lemma2 : def (in-codomain (OPwr (Ord (& A))) (_∩_ A)) (& t) + lemma2 not = ⊥-elim ( not (& t) ⟪ lemma3 , lemma6 ⟫ ) where + lemma6 : & t ≡ & (A ∩ * (& t)) + lemma6 = cong ( λ k → & k ) (==→o≡ (subst (λ k → t =h= (A ∩ k)) (sym *iso) ( ∩-≡ {t} {A} t→A ))) + + +extensionality0 : {A B : HOD } → ((z : HOD) → (A ∋ z) ⇔ (B ∋ z)) → A =h= B +eq→ (extensionality0 {A} {B} eq ) {x} d = odef-iso {A} {B} (sym &iso) (proj1 (eq (* x))) d +eq← (extensionality0 {A} {B} eq ) {x} d = odef-iso {B} {A} (sym &iso) (proj2 (eq (* x))) d + +extensionality : {A B w : HOD } → ((z : HOD ) → (A ∋ z) ⇔ (B ∋ z)) → (w ∋ A) ⇔ (w ∋ B) +proj1 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) ( ==→o≡ (extensionality0 {A} {B} eq) ) d +proj2 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) (sym ( ==→o≡ (extensionality0 {A} {B} eq) )) d + +infinity∅ : infinite ∋ od∅ +infinity∅ = subst (λ k → odef infinite k ) lemma iφ where + lemma : o∅ ≡ & od∅ + lemma = let open ≡-Reasoning in begin + o∅ + ≡⟨ sym &iso ⟩ + & ( * o∅ ) + ≡⟨ cong ( λ k → & k ) o∅≡od∅ ⟩ + & od∅ + ∎ +infinity : (x : HOD) → infinite ∋ x → infinite ∋ Union (x , (x , x )) +infinity x lt = subst (λ k → odef infinite k ) lemma (isuc {& x} lt) where + lemma : & (Union (* (& x) , (* (& x) , * (& x)))) + ≡ & (Union (x , (x , x))) + lemma = cong (λ k → & (Union ( k , ( k , k ) ))) *iso + +isZF : IsZF (HOD ) _∋_ _=h=_ od∅ _,_ Union Power Select Replace infinite +isZF = record { + isEquivalence = record { refl = ==-refl ; sym = ==-sym; trans = ==-trans } + ; pair→ = pair→ + ; pair← = pair← + ; union→ = union→ + ; union← = union← + ; empty = empty + ; power→ = power→ + ; power← = power← + ; extensionality = λ {A} {B} {w} → extensionality {A} {B} {w} + ; ε-induction = ε-induction + ; infinity∅ = infinity∅ + ; infinity = infinity + ; selection = λ {X} {ψ} {y} → selection {X} {ψ} {y} + ; replacement← = replacement← + ; replacement→ = λ {ψ} → replacement→ {ψ} + } + +HOD→ZF : ZF +HOD→ZF = record { + ZFSet = HOD + ; _∋_ = _∋_ + ; _≈_ = _=h=_ + ; ∅ = od∅ + ; _,_ = _,_ + ; Union = Union + ; Power = Power + ; Select = Select + ; Replace = Replace + ; infinite = infinite + ; isZF = isZF + } + +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ODC.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,128 @@ +{-# OPTIONS --allow-unsolved-metas #-} +open import Level +open import Ordinals +module ODC {n : Level } (O : Ordinals {n} ) where + +open import zf +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +open import Relation.Binary.PropositionalEquality +open import Data.Nat.Properties +open import Data.Empty +open import Relation.Nullary +open import Relation.Binary +open import Relation.Binary.Core + +import OrdUtil +open import logic +open import nat +import OD +import ODUtil + +open inOrdinal O +open OD O +open OD.OD +open OD._==_ +open ODAxiom odAxiom +open ODUtil O + +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +open OrdUtil O + + +open HOD + +open _∧_ + +postulate + -- mimimul and x∋minimal is an Axiom of choice + minimal : (x : HOD ) → ¬ (x =h= od∅ )→ HOD + -- this should be ¬ (x =h= od∅ )→ ∃ ox → x ∋ Ord ox ( minimum of x ) + x∋minimal : (x : HOD ) → ( ne : ¬ (x =h= od∅ ) ) → odef x ( & ( minimal x ne ) ) + -- minimality (may proved by ε-induction with LEM) + minimal-1 : (x : HOD ) → ( ne : ¬ (x =h= od∅ ) ) → (y : HOD ) → ¬ ( odef (minimal x ne) (& y)) ∧ (odef x (& y) ) + + +-- +-- Axiom of choice in intutionistic logic implies the exclude middle +-- https://plato.stanford.edu/entries/axiom-choice/#AxiChoLog +-- + +pred-od : ( p : Set n ) → HOD +pred-od p = record { od = record { def = λ x → (x ≡ o∅) ∧ p } ; + odmax = osuc o∅; <odmax = λ x → subst (λ k → k o< osuc o∅) (sym (proj1 x)) <-osuc } + +ppp : { p : Set n } { a : HOD } → pred-od p ∋ a → p +ppp {p} {a} d = proj2 d + +p∨¬p : ( p : Set n ) → p ∨ ( ¬ p ) -- assuming axiom of choice +p∨¬p p with is-o∅ ( & (pred-od p )) +p∨¬p p | yes eq = case2 (¬p eq) where + ps = pred-od p + eqo∅ : ps =h= od∅ → & ps ≡ o∅ + eqo∅ eq = subst (λ k → & ps ≡ k) ord-od∅ ( cong (λ k → & k ) (==→o≡ eq)) + lemma : ps =h= od∅ → p → ⊥ + lemma eq p0 = ¬x<0 {& ps} (eq→ eq record { proj1 = eqo∅ eq ; proj2 = p0 } ) + ¬p : (& ps ≡ o∅) → p → ⊥ + ¬p eq = lemma ( subst₂ (λ j k → j =h= k ) *iso o∅≡od∅ ( o≡→== eq )) +p∨¬p p | no ¬p = case1 (ppp {p} {minimal ps (λ eq → ¬p (eqo∅ eq))} lemma) where + ps = pred-od p + eqo∅ : ps =h= od∅ → & ps ≡ o∅ + eqo∅ eq = subst (λ k → & ps ≡ k) ord-od∅ ( cong (λ k → & k ) (==→o≡ eq)) + lemma : ps ∋ minimal ps (λ eq → ¬p (eqo∅ eq)) + lemma = x∋minimal ps (λ eq → ¬p (eqo∅ eq)) + +decp : ( p : Set n ) → Dec p -- assuming axiom of choice +decp p with p∨¬p p +decp p | case1 x = yes x +decp p | case2 x = no x + +∋-p : (A x : HOD ) → Dec ( A ∋ x ) +∋-p A x with p∨¬p ( A ∋ x ) -- LEM +∋-p A x | case1 t = yes t +∋-p A x | case2 t = no (λ x → t x) + +double-neg-eilm : {A : Set n} → ¬ ¬ A → A -- we don't have this in intutionistic logic +double-neg-eilm {A} notnot with decp A -- assuming axiom of choice +... | yes p = p +... | no ¬p = ⊥-elim ( notnot ¬p ) + +open _⊆_ + +power→⊆ : ( A t : HOD) → Power A ∋ t → t ⊆ A +power→⊆ A t PA∋t = record { incl = λ {x} t∋x → double-neg-eilm (t1 t∋x) } where + t1 : {x : HOD } → t ∋ x → ¬ ¬ (A ∋ x) + t1 = power→ A t PA∋t + +OrdP : ( x : Ordinal ) ( y : HOD ) → Dec ( Ord x ∋ y ) +OrdP x y with trio< x (& y) +OrdP x y | tri< a ¬b ¬c = no ¬c +OrdP x y | tri≈ ¬a refl ¬c = no ( o<¬≡ refl ) +OrdP x y | tri> ¬a ¬b c = yes c + +open import zfc + +HOD→ZFC : ZFC +HOD→ZFC = record { + ZFSet = HOD + ; _∋_ = _∋_ + ; _≈_ = _=h=_ + ; ∅ = od∅ + ; Select = Select + ; isZFC = isZFC + } where + -- infixr 200 _∈_ + -- infixr 230 _∩_ _∪_ + isZFC : IsZFC (HOD ) _∋_ _=h=_ od∅ Select + isZFC = record { + choice-func = choice-func ; + choice = choice + } where + -- Axiom of choice ( is equivalent to the existence of minimal in our case ) + -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] + choice-func : (X : HOD ) → {x : HOD } → ¬ ( x =h= od∅ ) → ( X ∋ x ) → HOD + choice-func X {x} not X∋x = minimal x not + choice : (X : HOD ) → {A : HOD } → ( X∋A : X ∋ A ) → (not : ¬ ( A =h= od∅ )) → A ∋ choice-func X not X∋A + choice X {A} X∋A not = x∋minimal A not +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ODUtil.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,179 @@ +{-# OPTIONS --allow-unsolved-metas #-} +open import Level +open import Ordinals +module ODUtil {n : Level } (O : Ordinals {n} ) where + +open import zf +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +open import Relation.Binary.PropositionalEquality hiding ( [_] ) +open import Data.Nat.Properties +open import Data.Empty +open import Relation.Nullary +open import Relation.Binary hiding ( _⇔_ ) + +open import logic +open import nat + +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +import OrdUtil +open OrdUtil O + +import OD +open OD O +open OD.OD +open ODAxiom odAxiom + +open HOD +open _⊆_ +open _∧_ +open _==_ + +cseq : HOD → HOD +cseq x = record { od = record { def = λ y → odef x (osuc y) } ; odmax = osuc (odmax x) ; <odmax = lemma } where + lemma : {y : Ordinal} → def (od x) (osuc y) → y o< osuc (odmax x) + lemma {y} lt = ordtrans <-osuc (ordtrans (<odmax x lt) <-osuc ) + + +pair-xx<xy : {x y : HOD} → & (x , x) o< osuc (& (x , y) ) +pair-xx<xy {x} {y} = ⊆→o≤ lemma where + lemma : {z : Ordinal} → def (od (x , x)) z → def (od (x , y)) z + lemma {z} (case1 refl) = case1 refl + lemma {z} (case2 refl) = case1 refl + +pair-<xy : {x y : HOD} → {n : Ordinal} → & x o< next n → & y o< next n → & (x , y) o< next n +pair-<xy {x} {y} {o} x<nn y<nn with trio< (& x) (& y) | inspect (omax (& x)) (& y) +... | tri< a ¬b ¬c | record { eq = eq1 } = next< (subst (λ k → k o< next o ) (sym eq1) (osuc<nx y<nn)) ho< +... | tri> ¬a ¬b c | record { eq = eq1 } = next< (subst (λ k → k o< next o ) (sym eq1) (osuc<nx x<nn)) ho< +... | tri≈ ¬a b ¬c | record { eq = eq1 } = next< (subst (λ k → k o< next o ) (omax≡ _ _ b) (subst (λ k → osuc k o< next o) b (osuc<nx x<nn))) ho< + +-- another form of infinite +-- pair-ord< : {x : Ordinal } → Set n +pair-ord< : {x : HOD } → ( {y : HOD } → & y o< next (odmax y) ) → & ( x , x ) o< next (& x) +pair-ord< {x} ho< = subst (λ k → & (x , x) o< k ) lemmab0 lemmab1 where + lemmab0 : next (odmax (x , x)) ≡ next (& x) + lemmab0 = trans (cong (λ k → next k) (omxx _)) (sym nexto≡) + lemmab1 : & (x , x) o< next ( odmax (x , x)) + lemmab1 = ho< + +trans-⊆ : { A B C : HOD} → A ⊆ B → B ⊆ C → A ⊆ C +trans-⊆ A⊆B B⊆C = record { incl = λ x → incl B⊆C (incl A⊆B x) } + +refl-⊆ : {A : HOD} → A ⊆ A +refl-⊆ {A} = record { incl = λ x → x } + +od⊆→o≤ : {x y : HOD } → x ⊆ y → & x o< osuc (& y) +od⊆→o≤ {x} {y} lt = ⊆→o≤ {x} {y} (λ {z} x>z → subst (λ k → def (od y) k ) &iso (incl lt (d→∋ x x>z))) + +subset-lemma : {A x : HOD } → ( {y : HOD } → x ∋ y → (A ∩ x ) ∋ y ) ⇔ ( x ⊆ A ) +subset-lemma {A} {x} = record { + proj1 = λ lt → record { incl = λ x∋z → proj1 (lt x∋z) } + ; proj2 = λ x⊆A lt → ⟪ incl x⊆A lt , lt ⟫ + } + + +ω<next-o∅ : {y : Ordinal} → infinite-d y → y o< next o∅ +ω<next-o∅ {y} lt = <odmax infinite lt + +nat→ω : Nat → HOD +nat→ω Zero = od∅ +nat→ω (Suc y) = Union (nat→ω y , (nat→ω y , nat→ω y)) + +ω→nato : {y : Ordinal} → infinite-d y → Nat +ω→nato iφ = Zero +ω→nato (isuc lt) = Suc (ω→nato lt) + +ω→nat : (n : HOD) → infinite ∋ n → Nat +ω→nat n = ω→nato + +ω∋nat→ω : {n : Nat} → def (od infinite) (& (nat→ω n)) +ω∋nat→ω {Zero} = subst (λ k → def (od infinite) k) (sym ord-od∅) iφ +ω∋nat→ω {Suc n} = subst (λ k → def (od infinite) k) lemma (isuc ( ω∋nat→ω {n})) where + lemma : & (Union (* (& (nat→ω n)) , (* (& (nat→ω n)) , * (& (nat→ω n))))) ≡ & (nat→ω (Suc n)) + lemma = subst (λ k → & (Union (k , ( k , k ))) ≡ & (nat→ω (Suc n))) (sym *iso) refl + +pair1 : { x y : HOD } → (x , y ) ∋ x +pair1 = case1 refl + +pair2 : { x y : HOD } → (x , y ) ∋ y +pair2 = case2 refl + +single : {x y : HOD } → (x , x ) ∋ y → x ≡ y +single (case1 eq) = ==→o≡ ( ord→== (sym eq) ) +single (case2 eq) = ==→o≡ ( ord→== (sym eq) ) + +open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) +-- postulate f-extensionality : { n m : Level} → HE.Extensionality n m + +ω-prev-eq1 : {x y : Ordinal} → & (Union (* y , (* y , * y))) ≡ & (Union (* x , (* x , * x))) → ¬ (x o< y) +ω-prev-eq1 {x} {y} eq x<y = eq→ (ord→== eq) {& (* y)} (λ not2 → not2 (& (* y , * y)) + ⟪ case2 refl , subst (λ k → odef k (& (* y))) (sym *iso) (case1 refl)⟫ ) (λ u → lemma u ) where + lemma : (u : Ordinal) → ¬ def (od (* x , (* x , * x))) u ∧ def (od (* u)) (& (* y)) + lemma u t with proj1 t + lemma u t | case1 u=x = o<> (c<→o< {* y} {* u} (proj2 t)) (subst₂ (λ j k → j o< k ) + (trans (sym &iso) (trans (sym u=x) (sym &iso)) ) (sym &iso) x<y ) -- x ≡ & (* u) + lemma u t | case2 u=xx = o<¬≡ (lemma1 (subst (λ k → odef k (& (* y)) ) (trans (cong (λ k → * k ) u=xx) *iso ) (proj2 t))) x<y where + lemma1 : {x y : Ordinal } → (* x , * x ) ∋ * y → x ≡ y -- y = x ∈ ( x , x ) = u + lemma1 (case1 eq) = subst₂ (λ j k → j ≡ k ) &iso &iso (sym eq) + lemma1 (case2 eq) = subst₂ (λ j k → j ≡ k ) &iso &iso (sym eq) + +ω-prev-eq : {x y : Ordinal} → & (Union (* y , (* y , * y))) ≡ & (Union (* x , (* x , * x))) → x ≡ y +ω-prev-eq {x} {y} eq with trio< x y +ω-prev-eq {x} {y} eq | tri< a ¬b ¬c = ⊥-elim (ω-prev-eq1 eq a) +ω-prev-eq {x} {y} eq | tri≈ ¬a b ¬c = b +ω-prev-eq {x} {y} eq | tri> ¬a ¬b c = ⊥-elim (ω-prev-eq1 (sym eq) c) + +ω-∈s : (x : HOD) → Union ( x , (x , x)) ∋ x +ω-∈s x not = not (& (x , x)) ⟪ case2 refl , subst (λ k → odef k (& x) ) (sym *iso) (case1 refl) ⟫ + +ωs≠0 : (x : HOD) → ¬ ( Union ( x , (x , x)) ≡ od∅ ) +ωs≠0 y eq = ⊥-elim ( ¬x<0 (subst (λ k → & y o< k ) ord-od∅ (c<→o< (subst (λ k → odef k (& y )) eq (ω-∈s y) ))) ) + +nat→ω-iso : {i : HOD} → (lt : infinite ∋ i ) → nat→ω ( ω→nat i lt ) ≡ i +nat→ω-iso {i} = ε-induction {λ i → (lt : infinite ∋ i ) → nat→ω ( ω→nat i lt ) ≡ i } ind i where + ind : {x : HOD} → ({y : HOD} → x ∋ y → (lt : infinite ∋ y) → nat→ω (ω→nat y lt) ≡ y) → + (lt : infinite ∋ x) → nat→ω (ω→nat x lt) ≡ x + ind {x} prev lt = ind1 lt *iso where + ind1 : {ox : Ordinal } → (ltd : infinite-d ox ) → * ox ≡ x → nat→ω (ω→nato ltd) ≡ x + ind1 {o∅} iφ refl = sym o∅≡od∅ + ind1 (isuc {x₁} ltd) ox=x = begin + nat→ω (ω→nato (isuc ltd) ) + ≡⟨⟩ + Union (nat→ω (ω→nato ltd) , (nat→ω (ω→nato ltd) , nat→ω (ω→nato ltd))) + ≡⟨ cong (λ k → Union (k , (k , k ))) lemma ⟩ + Union (* x₁ , (* x₁ , * x₁)) + ≡⟨ trans ( sym *iso) ox=x ⟩ + x + ∎ where + open ≡-Reasoning + lemma0 : x ∋ * x₁ + lemma0 = subst (λ k → odef k (& (* x₁))) (trans (sym *iso) ox=x) (λ not → not + (& (* x₁ , * x₁)) ⟪ pair2 , subst (λ k → odef k (& (* x₁))) (sym *iso) pair1 ⟫ ) + lemma1 : infinite ∋ * x₁ + lemma1 = subst (λ k → odef infinite k) (sym &iso) ltd + lemma3 : {x y : Ordinal} → (ltd : infinite-d x ) (ltd1 : infinite-d y ) → y ≡ x → ltd ≅ ltd1 + lemma3 iφ iφ refl = HE.refl + lemma3 iφ (isuc {y} ltd1) eq = ⊥-elim ( ¬x<0 (subst₂ (λ j k → j o< k ) &iso eq (c<→o< (ω-∈s (* y)) ))) + lemma3 (isuc {y} ltd) iφ eq = ⊥-elim ( ¬x<0 (subst₂ (λ j k → j o< k ) &iso (sym eq) (c<→o< (ω-∈s (* y)) ))) + lemma3 (isuc {x} ltd) (isuc {y} ltd1) eq with lemma3 ltd ltd1 (ω-prev-eq (sym eq)) + ... | t = HE.cong₂ (λ j k → isuc {j} k ) (HE.≡-to-≅ (ω-prev-eq eq)) t + lemma2 : {x y : Ordinal} → (ltd : infinite-d x ) (ltd1 : infinite-d y ) → y ≡ x → ω→nato ltd ≡ ω→nato ltd1 + lemma2 {x} {y} ltd ltd1 eq = lemma6 eq (lemma3 {x} {y} ltd ltd1 eq) where + lemma6 : {x y : Ordinal} → {ltd : infinite-d x } {ltd1 : infinite-d y } → y ≡ x → ltd ≅ ltd1 → ω→nato ltd ≡ ω→nato ltd1 + lemma6 refl HE.refl = refl + lemma : nat→ω (ω→nato ltd) ≡ * x₁ + lemma = trans (cong (λ k → nat→ω k) (lemma2 {x₁} {_} ltd (subst (λ k → infinite-d k ) (sym &iso) ltd) &iso ) ) ( prev {* x₁} lemma0 lemma1 ) + +ω→nat-iso : {i : Nat} → ω→nat ( nat→ω i ) (ω∋nat→ω {i}) ≡ i +ω→nat-iso {i} = lemma i (ω∋nat→ω {i}) *iso where + lemma : {x : Ordinal } → ( i : Nat ) → (ltd : infinite-d x ) → * x ≡ nat→ω i → ω→nato ltd ≡ i + lemma {x} Zero iφ eq = refl + lemma {x} (Suc i) iφ eq = ⊥-elim ( ωs≠0 (nat→ω i) (trans (sym eq) o∅≡od∅ )) -- Union (nat→ω i , (nat→ω i , nat→ω i)) ≡ od∅ + lemma Zero (isuc {x} ltd) eq = ⊥-elim ( ωs≠0 (* x) (subst (λ k → k ≡ od∅ ) *iso eq )) + lemma (Suc i) (isuc {x} ltd) eq = cong (λ k → Suc k ) (lemma i ltd (lemma1 eq) ) where -- * x ≡ nat→ω i + lemma1 : * (& (Union (* x , (* x , * x)))) ≡ Union (nat→ω i , (nat→ω i , nat→ω i)) → * x ≡ nat→ω i + lemma1 eq = subst (λ k → * x ≡ k ) *iso (cong (λ k → * k) + ( ω-prev-eq (subst (λ k → _ ≡ k ) &iso (cong (λ k → & k ) (sym + (subst (λ k → _ ≡ Union ( k , ( k , k ))) (sym *iso ) eq )))))) +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/OPair.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,213 @@ +{-# OPTIONS --allow-unsolved-metas #-} + +open import Level +open import Ordinals +module OPair {n : Level } (O : Ordinals {n}) where + +open import zf +open import logic +import OD +import ODUtil +import OrdUtil + +open import Relation.Nullary +open import Relation.Binary +open import Data.Empty +open import Relation.Binary +open import Relation.Binary.Core +open import Relation.Binary.PropositionalEquality +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) + +open OD O +open OD.OD +open OD.HOD +open ODAxiom odAxiom + +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +open OrdUtil O +open ODUtil O + +open _∧_ +open _∨_ +open Bool + +open _==_ + +<_,_> : (x y : HOD) → HOD +< x , y > = (x , x ) , (x , y ) + +exg-pair : { x y : HOD } → (x , y ) =h= ( y , x ) +exg-pair {x} {y} = record { eq→ = left ; eq← = right } where + left : {z : Ordinal} → odef (x , y) z → odef (y , x) z + left (case1 t) = case2 t + left (case2 t) = case1 t + right : {z : Ordinal} → odef (y , x) z → odef (x , y) z + right (case1 t) = case2 t + right (case2 t) = case1 t + +ord≡→≡ : { x y : HOD } → & x ≡ & y → x ≡ y +ord≡→≡ eq = subst₂ (λ j k → j ≡ k ) *iso *iso ( cong ( λ k → * k ) eq ) + +od≡→≡ : { x y : Ordinal } → * x ≡ * y → x ≡ y +od≡→≡ eq = subst₂ (λ j k → j ≡ k ) &iso &iso ( cong ( λ k → & k ) eq ) + +eq-prod : { x x' y y' : HOD } → x ≡ x' → y ≡ y' → < x , y > ≡ < x' , y' > +eq-prod refl refl = refl + +xx=zy→x=y : {x y z : HOD } → ( x , x ) =h= ( z , y ) → x ≡ y +xx=zy→x=y {x} {y} eq with trio< (& x) (& y) +xx=zy→x=y {x} {y} eq | tri< a ¬b ¬c with eq← eq {& y} (case2 refl) +xx=zy→x=y {x} {y} eq | tri< a ¬b ¬c | case1 s = ⊥-elim ( o<¬≡ (sym s) a ) +xx=zy→x=y {x} {y} eq | tri< a ¬b ¬c | case2 s = ⊥-elim ( o<¬≡ (sym s) a ) +xx=zy→x=y {x} {y} eq | tri≈ ¬a b ¬c = ord≡→≡ b +xx=zy→x=y {x} {y} eq | tri> ¬a ¬b c with eq← eq {& y} (case2 refl) +xx=zy→x=y {x} {y} eq | tri> ¬a ¬b c | case1 s = ⊥-elim ( o<¬≡ s c ) +xx=zy→x=y {x} {y} eq | tri> ¬a ¬b c | case2 s = ⊥-elim ( o<¬≡ s c ) + +prod-eq : { x x' y y' : HOD } → < x , y > =h= < x' , y' > → (x ≡ x' ) ∧ ( y ≡ y' ) +prod-eq {x} {x'} {y} {y'} eq = ⟪ lemmax , lemmay ⟫ where + lemma2 : {x y z : HOD } → ( x , x ) =h= ( z , y ) → z ≡ y + lemma2 {x} {y} {z} eq = trans (sym (xx=zy→x=y lemma3 )) ( xx=zy→x=y eq ) where + lemma3 : ( x , x ) =h= ( y , z ) + lemma3 = ==-trans eq exg-pair + lemma1 : {x y : HOD } → ( x , x ) =h= ( y , y ) → x ≡ y + lemma1 {x} {y} eq with eq← eq {& y} (case2 refl) + lemma1 {x} {y} eq | case1 s = ord≡→≡ (sym s) + lemma1 {x} {y} eq | case2 s = ord≡→≡ (sym s) + lemma4 : {x y z : HOD } → ( x , y ) =h= ( x , z ) → y ≡ z + lemma4 {x} {y} {z} eq with eq← eq {& z} (case2 refl) + lemma4 {x} {y} {z} eq | case1 s with ord≡→≡ s -- x ≡ z + ... | refl with lemma2 (==-sym eq ) + ... | refl = refl + lemma4 {x} {y} {z} eq | case2 s = ord≡→≡ (sym s) -- y ≡ z + lemmax : x ≡ x' + lemmax with eq→ eq {& (x , x)} (case1 refl) + lemmax | case1 s = lemma1 (ord→== s ) -- (x,x)≡(x',x') + lemmax | case2 s with lemma2 (ord→== s ) -- (x,x)≡(x',y') with x'≡y' + ... | refl = lemma1 (ord→== s ) + lemmay : y ≡ y' + lemmay with lemmax + ... | refl with lemma4 eq -- with (x,y)≡(x,y') + ... | eq1 = lemma4 (ord→== (cong (λ k → & k ) eq1 )) + +-- +-- unlike ordered pair, ZFProduct is not a HOD + +data ord-pair : (p : Ordinal) → Set n where + pair : (x y : Ordinal ) → ord-pair ( & ( < * x , * y > ) ) + +ZFProduct : OD +ZFProduct = record { def = λ x → ord-pair x } + +-- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) +-- eq-pair : { x x' y y' : Ordinal } → x ≡ x' → y ≡ y' → pair x y ≅ pair x' y' +-- eq-pair refl refl = HE.refl + +pi1 : { p : Ordinal } → ord-pair p → Ordinal +pi1 ( pair x y) = x + +π1 : { p : HOD } → def ZFProduct (& p) → HOD +π1 lt = * (pi1 lt ) + +pi2 : { p : Ordinal } → ord-pair p → Ordinal +pi2 ( pair x y ) = y + +π2 : { p : HOD } → def ZFProduct (& p) → HOD +π2 lt = * (pi2 lt ) + +op-cons : { ox oy : Ordinal } → def ZFProduct (& ( < * ox , * oy > )) +op-cons {ox} {oy} = pair ox oy + +def-subst : {Z : OD } {X : Ordinal }{z : OD } {x : Ordinal }→ def Z X → Z ≡ z → X ≡ x → def z x +def-subst df refl refl = df + +p-cons : ( x y : HOD ) → def ZFProduct (& ( < x , y >)) +p-cons x y = def-subst {_} {_} {ZFProduct} {& (< x , y >)} (pair (& x) ( & y )) refl ( + let open ≡-Reasoning in begin + & < * (& x) , * (& y) > + ≡⟨ cong₂ (λ j k → & < j , k >) *iso *iso ⟩ + & < x , y > + ∎ ) + +op-iso : { op : Ordinal } → (q : ord-pair op ) → & < * (pi1 q) , * (pi2 q) > ≡ op +op-iso (pair ox oy) = refl + +p-iso : { x : HOD } → (p : def ZFProduct (& x) ) → < π1 p , π2 p > ≡ x +p-iso {x} p = ord≡→≡ (op-iso p) + +p-pi1 : { x y : HOD } → (p : def ZFProduct (& < x , y >) ) → π1 p ≡ x +p-pi1 {x} {y} p = proj1 ( prod-eq ( ord→== (op-iso p) )) + +p-pi2 : { x y : HOD } → (p : def ZFProduct (& < x , y >) ) → π2 p ≡ y +p-pi2 {x} {y} p = proj2 ( prod-eq ( ord→== (op-iso p))) + +ω-pair : {x y : HOD} → {m : Ordinal} → & x o< next m → & y o< next m → & (x , y) o< next m +ω-pair lx ly = next< (omax<nx lx ly ) ho< + +ω-opair : {x y : HOD} → {m : Ordinal} → & x o< next m → & y o< next m → & < x , y > o< next m +ω-opair {x} {y} {m} lx ly = lemma0 where + lemma0 : & < x , y > o< next m + lemma0 = osucprev (begin + osuc (& < x , y >) + <⟨ osuc<nx ho< ⟩ + next (omax (& (x , x)) (& (x , y))) + ≡⟨ cong (λ k → next k) (sym ( omax≤ _ _ pair-xx<xy )) ⟩ + next (osuc (& (x , y))) + ≡⟨ sym (nexto≡) ⟩ + next (& (x , y)) + ≤⟨ x<ny→≤next (ω-pair lx ly) ⟩ + next m + ∎ ) where + open o≤-Reasoning O + +_⊗_ : (A B : HOD) → HOD +A ⊗ B = Union ( Replace B (λ b → Replace A (λ a → < a , b > ) )) + +product→ : {A B a b : HOD} → A ∋ a → B ∋ b → ( A ⊗ B ) ∋ < a , b > +product→ {A} {B} {a} {b} A∋a B∋b = λ t → t (& (Replace A (λ a → < a , b >))) + ⟪ lemma1 , subst (λ k → odef k (& < a , b >)) (sym *iso) lemma2 ⟫ where + lemma1 : odef (Replace B (λ b₁ → Replace A (λ a₁ → < a₁ , b₁ >))) (& (Replace A (λ a₁ → < a₁ , b >))) + lemma1 = replacement← B b B∋b + lemma2 : odef (Replace A (λ a₁ → < a₁ , b >)) (& < a , b >) + lemma2 = replacement← A a A∋a + +x<nextA : {A x : HOD} → A ∋ x → & x o< next (odmax A) +x<nextA {A} {x} A∋x = ordtrans (c<→o< {x} {A} A∋x) ho< + +A<Bnext : {A B x : HOD} → & A o< & B → A ∋ x → & x o< next (odmax B) +A<Bnext {A} {B} {x} lt A∋x = osucprev (begin + osuc (& x) + <⟨ osucc (c<→o< A∋x) ⟩ + osuc (& A) + <⟨ osucc lt ⟩ + osuc (& B) + <⟨ osuc<nx ho< ⟩ + next (odmax B) + ∎ ) where open o≤-Reasoning O + +ZFP : (A B : HOD) → HOD +ZFP A B = record { od = record { def = λ x → ord-pair x ∧ ((p : ord-pair x ) → odef A (pi1 p) ∧ odef B (pi2 p) )} ; + odmax = omax (next (odmax A)) (next (odmax B)) ; <odmax = λ {y} px → lemma y px } + where + lemma : (y : Ordinal) → ( ord-pair y ∧ ((p : ord-pair y) → odef A (pi1 p) ∧ odef B (pi2 p))) → y o< omax (next (odmax A)) (next (odmax B)) + lemma y lt with proj1 lt + lemma p lt | pair x y with trio< (& A) (& B) + lemma p lt | pair x y | tri< a ¬b ¬c = ordtrans (ω-opair (A<Bnext a (subst (λ k → odef A k ) (sym &iso) + (proj1 (proj2 lt (pair x y))))) (lemma1 (proj2 (proj2 lt (pair x y))))) (omax-y _ _ ) where + lemma1 : odef B y → & (* y) o< next (HOD.odmax B) + lemma1 lt = x<nextA {B} (d→∋ B lt) + lemma p lt | pair x y | tri≈ ¬a b ¬c = ordtrans (ω-opair (x<nextA {A} (d→∋ A ((proj1 (proj2 lt (pair x y)))))) lemma2 ) (omax-x _ _ ) where + lemma2 : & (* y) o< next (HOD.odmax A) + lemma2 = ordtrans ( subst (λ k → & (* y) o< k ) (sym b) (c<→o< (d→∋ B ((proj2 (proj2 lt (pair x y))))))) ho< + lemma p lt | pair x y | tri> ¬a ¬b c = ordtrans (ω-opair (x<nextA {A} (d→∋ A ((proj1 (proj2 lt (pair x y)))))) + (A<Bnext c (subst (λ k → odef B k ) (sym &iso) (proj2 (proj2 lt (pair x y)))))) (omax-x _ _ ) + +ZFP⊆⊗ : {A B : HOD} {x : Ordinal} → odef (ZFP A B) x → odef (A ⊗ B) x +ZFP⊆⊗ {A} {B} {px} ⟪ (pair x y) , p2 ⟫ = product→ (d→∋ A (proj1 (p2 (pair x y) ))) (d→∋ B (proj2 (p2 (pair x y) ))) + +-- axiom of choice required +-- ⊗⊆ZFP : {A B x : HOD} → ( A ⊗ B ) ∋ x → def ZFProduct (& x) +-- ⊗⊆ZFP {A} {B} {x} lt = subst (λ k → ord-pair (& k )) {!!} op-cons +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/OrdUtil.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,287 @@ +open import Level +open import Ordinals +module OrdUtil {n : Level} (O : Ordinals {n} ) where + +open import logic +open import nat +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +open import Data.Empty +open import Relation.Binary.PropositionalEquality +open import Relation.Nullary +open import Relation.Binary hiding (_⇔_) + +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext + +o<-dom : { x y : Ordinal } → x o< y → Ordinal +o<-dom {x} _ = x + +o<-cod : { x y : Ordinal } → x o< y → Ordinal +o<-cod {_} {y} _ = y + +o<-subst : {Z X z x : Ordinal } → Z o< X → Z ≡ z → X ≡ x → z o< x +o<-subst df refl refl = df + +o<¬≡ : { ox oy : Ordinal } → ox ≡ oy → ox o< oy → ⊥ +o<¬≡ {ox} {oy} eq lt with trio< ox oy +o<¬≡ {ox} {oy} eq lt | tri< a ¬b ¬c = ¬b eq +o<¬≡ {ox} {oy} eq lt | tri≈ ¬a b ¬c = ¬a lt +o<¬≡ {ox} {oy} eq lt | tri> ¬a ¬b c = ¬b eq + +o<> : {x y : Ordinal } → y o< x → x o< y → ⊥ +o<> {ox} {oy} lt tl with trio< ox oy +o<> {ox} {oy} lt tl | tri< a ¬b ¬c = ¬c lt +o<> {ox} {oy} lt tl | tri≈ ¬a b ¬c = ¬a tl +o<> {ox} {oy} lt tl | tri> ¬a ¬b c = ¬a tl + +osuc-< : { x y : Ordinal } → y o< osuc x → x o< y → ⊥ +osuc-< {x} {y} y<ox x<y with osuc-≡< y<ox +osuc-< {x} {y} y<ox x<y | case1 refl = o<¬≡ refl x<y +osuc-< {x} {y} y<ox x<y | case2 y<x = o<> x<y y<x + +osucc : {ox oy : Ordinal } → oy o< ox → osuc oy o< osuc ox +---- y < osuc y < x < osuc x +---- y < osuc y = x < osuc x +---- y < osuc y > x < osuc x -> y = x ∨ x < y → ⊥ +osucc {ox} {oy} oy<ox with trio< (osuc oy) ox +osucc {ox} {oy} oy<ox | tri< a ¬b ¬c = ordtrans a <-osuc +osucc {ox} {oy} oy<ox | tri≈ ¬a refl ¬c = <-osuc +osucc {ox} {oy} oy<ox | tri> ¬a ¬b c with osuc-≡< c +osucc {ox} {oy} oy<ox | tri> ¬a ¬b c | case1 eq = ⊥-elim (o<¬≡ (sym eq) oy<ox) +osucc {ox} {oy} oy<ox | tri> ¬a ¬b c | case2 lt = ⊥-elim (o<> lt oy<ox) + +osucprev : {ox oy : Ordinal } → osuc oy o< osuc ox → oy o< ox +osucprev {ox} {oy} oy<ox with trio< oy ox +osucprev {ox} {oy} oy<ox | tri< a ¬b ¬c = a +osucprev {ox} {oy} oy<ox | tri≈ ¬a b ¬c = ⊥-elim (o<¬≡ (cong (λ k → osuc k) b) oy<ox ) +osucprev {ox} {oy} oy<ox | tri> ¬a ¬b c = ⊥-elim (o<> (osucc c) oy<ox ) + +open _∧_ + +osuc2 : ( x y : Ordinal ) → ( osuc x o< osuc (osuc y )) ⇔ (x o< osuc y) +proj2 (osuc2 x y) lt = osucc lt +proj1 (osuc2 x y) ox<ooy with osuc-≡< ox<ooy +proj1 (osuc2 x y) ox<ooy | case1 ox=oy = o<-subst <-osuc refl ox=oy +proj1 (osuc2 x y) ox<ooy | case2 ox<oy = ordtrans <-osuc ox<oy + +_o≤_ : Ordinal → Ordinal → Set n +a o≤ b = a o< osuc b -- (a ≡ b) ∨ ( a o< b ) + + +xo<ab : {oa ob : Ordinal } → ( {ox : Ordinal } → ox o< oa → ox o< ob ) → oa o< osuc ob +xo<ab {oa} {ob} a→b with trio< oa ob +xo<ab {oa} {ob} a→b | tri< a ¬b ¬c = ordtrans a <-osuc +xo<ab {oa} {ob} a→b | tri≈ ¬a refl ¬c = <-osuc +xo<ab {oa} {ob} a→b | tri> ¬a ¬b c = ⊥-elim ( o<¬≡ refl (a→b c ) ) + +maxα : Ordinal → Ordinal → Ordinal +maxα x y with trio< x y +maxα x y | tri< a ¬b ¬c = y +maxα x y | tri> ¬a ¬b c = x +maxα x y | tri≈ ¬a refl ¬c = x + +omin : Ordinal → Ordinal → Ordinal +omin x y with trio< x y +omin x y | tri< a ¬b ¬c = x +omin x y | tri> ¬a ¬b c = y +omin x y | tri≈ ¬a refl ¬c = x + +min1 : {x y z : Ordinal } → z o< x → z o< y → z o< omin x y +min1 {x} {y} {z} z<x z<y with trio< x y +min1 {x} {y} {z} z<x z<y | tri< a ¬b ¬c = z<x +min1 {x} {y} {z} z<x z<y | tri≈ ¬a refl ¬c = z<x +min1 {x} {y} {z} z<x z<y | tri> ¬a ¬b c = z<y + +-- +-- max ( osuc x , osuc y ) +-- + +omax : ( x y : Ordinal ) → Ordinal +omax x y with trio< x y +omax x y | tri< a ¬b ¬c = osuc y +omax x y | tri> ¬a ¬b c = osuc x +omax x y | tri≈ ¬a refl ¬c = osuc x + +omax< : ( x y : Ordinal ) → x o< y → osuc y ≡ omax x y +omax< x y lt with trio< x y +omax< x y lt | tri< a ¬b ¬c = refl +omax< x y lt | tri≈ ¬a b ¬c = ⊥-elim (¬a lt ) +omax< x y lt | tri> ¬a ¬b c = ⊥-elim (¬a lt ) + +omax≤ : ( x y : Ordinal ) → x o≤ y → osuc y ≡ omax x y +omax≤ x y le with trio< x y +omax≤ x y le | tri< a ¬b ¬c = refl +omax≤ x y le | tri≈ ¬a refl ¬c = refl +omax≤ x y le | tri> ¬a ¬b c with osuc-≡< le +omax≤ x y le | tri> ¬a ¬b c | case1 eq = ⊥-elim (¬b eq) +omax≤ x y le | tri> ¬a ¬b c | case2 x<y = ⊥-elim (¬a x<y) + +omax≡ : ( x y : Ordinal ) → x ≡ y → osuc y ≡ omax x y +omax≡ x y eq with trio< x y +omax≡ x y eq | tri< a ¬b ¬c = ⊥-elim (¬b eq ) +omax≡ x y eq | tri≈ ¬a refl ¬c = refl +omax≡ x y eq | tri> ¬a ¬b c = ⊥-elim (¬b eq ) + +omax-x : ( x y : Ordinal ) → x o< omax x y +omax-x x y with trio< x y +omax-x x y | tri< a ¬b ¬c = ordtrans a <-osuc +omax-x x y | tri> ¬a ¬b c = <-osuc +omax-x x y | tri≈ ¬a refl ¬c = <-osuc + +omax-y : ( x y : Ordinal ) → y o< omax x y +omax-y x y with trio< x y +omax-y x y | tri< a ¬b ¬c = <-osuc +omax-y x y | tri> ¬a ¬b c = ordtrans c <-osuc +omax-y x y | tri≈ ¬a refl ¬c = <-osuc + +omxx : ( x : Ordinal ) → omax x x ≡ osuc x +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 + +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 )) + +open _∧_ + +o≤-refl : { i j : Ordinal } → i ≡ j → i o≤ j +o≤-refl {i} {j} eq = subst (λ k → i o< osuc k ) eq <-osuc +OrdTrans : Transitive _o≤_ +OrdTrans a≤b b≤c with osuc-≡< a≤b | osuc-≡< b≤c +OrdTrans a≤b b≤c | case1 refl | case1 refl = <-osuc +OrdTrans a≤b b≤c | case1 refl | case2 a≤c = ordtrans a≤c <-osuc +OrdTrans a≤b b≤c | case2 a≤c | case1 refl = ordtrans a≤c <-osuc +OrdTrans a≤b b≤c | case2 a<b | case2 b<c = ordtrans (ordtrans a<b b<c) <-osuc + +OrdPreorder : Preorder n n n +OrdPreorder = record { Carrier = Ordinal + ; _≈_ = _≡_ + ; _∼_ = _o≤_ + ; isPreorder = record { + isEquivalence = record { refl = refl ; sym = sym ; trans = trans } + ; reflexive = o≤-refl + ; trans = OrdTrans + } + } + +FExists : {m l : Level} → ( ψ : Ordinal → Set m ) + → {p : Set l} ( P : { y : Ordinal } → ψ y → ¬ p ) + → (exists : ¬ (∀ y → ¬ ( ψ y ) )) + → ¬ p +FExists {m} {l} ψ {p} P = contra-position ( λ p y ψy → P {y} ψy p ) + +nexto∅ : {x : Ordinal} → o∅ o< next x +nexto∅ {x} with trio< o∅ x +nexto∅ {x} | tri< a ¬b ¬c = ordtrans a x<nx +nexto∅ {x} | tri≈ ¬a b ¬c = subst (λ k → k o< next x) (sym b) x<nx +nexto∅ {x} | tri> ¬a ¬b c = ⊥-elim ( ¬x<0 c ) + +next< : {x y z : Ordinal} → x o< next z → y o< next x → y o< next z +next< {x} {y} {z} x<nz y<nx with trio< y (next z) +next< {x} {y} {z} x<nz y<nx | tri< a ¬b ¬c = a +next< {x} {y} {z} x<nz y<nx | tri≈ ¬a b ¬c = ⊥-elim (¬nx<nx x<nz (subst (λ k → k o< next x) b y<nx) + (λ w nz=ow → o<¬≡ nz=ow (subst₂ (λ j k → j o< k ) (sym nz=ow) nz=ow (osuc<nx (subst (λ k → w o< k ) (sym nz=ow) <-osuc) )))) +next< {x} {y} {z} x<nz y<nx | tri> ¬a ¬b c = ⊥-elim (¬nx<nx x<nz (ordtrans c y<nx ) + (λ w nz=ow → o<¬≡ (sym nz=ow) (osuc<nx (subst (λ k → w o< k ) (sym nz=ow) <-osuc )))) + +osuc< : {x y : Ordinal} → osuc x ≡ y → x o< y +osuc< {x} {y} refl = <-osuc + +nexto=n : {x y : Ordinal} → x o< next (osuc y) → x o< next y +nexto=n {x} {y} x<noy = next< (osuc<nx x<nx) x<noy + +nexto≡ : {x : Ordinal} → next x ≡ next (osuc x) +nexto≡ {x} with trio< (next x) (next (osuc x) ) +-- next x o< next (osuc x ) -> osuc x o< next x o< next (osuc x) -> next x ≡ osuc z -> z o o< next x -> osuc z o< next x -> next x o< next x +nexto≡ {x} | tri< a ¬b ¬c = ⊥-elim (¬nx<nx (osuc<nx x<nx ) a + (λ z eq → o<¬≡ (sym eq) (osuc<nx (osuc< (sym eq))))) +nexto≡ {x} | tri≈ ¬a b ¬c = b +-- next (osuc x) o< next x -> osuc x o< next (osuc x) o< next x -> next (osuc x) ≡ osuc z -> z o o< next (osuc x) ... +nexto≡ {x} | tri> ¬a ¬b c = ⊥-elim (¬nx<nx (ordtrans <-osuc x<nx) c + (λ z eq → o<¬≡ (sym eq) (osuc<nx (osuc< (sym eq))))) + +next-is-limit : {x y : Ordinal} → ¬ (next x ≡ osuc y) +next-is-limit {x} {y} eq = o<¬≡ (sym eq) (osuc<nx y<nx) where + y<nx : y o< next x + y<nx = osuc< (sym eq) + +omax<next : {x y : Ordinal} → x o< y → omax x y o< next y +omax<next {x} {y} x<y = subst (λ k → k o< next y ) (omax< _ _ x<y ) (osuc<nx x<nx) + +x<ny→≡next : {x y : Ordinal} → x o< y → y o< next x → next x ≡ next y +x<ny→≡next {x} {y} x<y y<nx with trio< (next x) (next y) +x<ny→≡next {x} {y} x<y y<nx | tri< a ¬b ¬c = -- x < y < next x < next y ∧ next x = osuc z + ⊥-elim ( ¬nx<nx y<nx a (λ z eq → o<¬≡ (sym eq) (osuc<nx (subst (λ k → z o< k ) (sym eq) <-osuc )))) +x<ny→≡next {x} {y} x<y y<nx | tri≈ ¬a b ¬c = b +x<ny→≡next {x} {y} x<y y<nx | tri> ¬a ¬b c = -- x < y < next y < next x + ⊥-elim ( ¬nx<nx (ordtrans x<y x<nx) c (λ z eq → o<¬≡ (sym eq) (osuc<nx (subst (λ k → z o< k ) (sym eq) <-osuc )))) + +≤next : {x y : Ordinal} → x o≤ y → next x o≤ next y +≤next {x} {y} x≤y with trio< (next x) y +≤next {x} {y} x≤y | tri< a ¬b ¬c = ordtrans a (ordtrans x<nx <-osuc ) +≤next {x} {y} x≤y | tri≈ ¬a refl ¬c = (ordtrans x<nx <-osuc ) +≤next {x} {y} x≤y | tri> ¬a ¬b c with osuc-≡< x≤y +≤next {x} {y} x≤y | tri> ¬a ¬b c | case1 refl = o≤-refl refl -- x = y < next x +≤next {x} {y} x≤y | tri> ¬a ¬b c | case2 x<y = o≤-refl (x<ny→≡next x<y c) -- x ≤ y < next x + +x<ny→≤next : {x y : Ordinal} → x o< next y → next x o≤ next y +x<ny→≤next {x} {y} x<ny with trio< x y +x<ny→≤next {x} {y} x<ny | tri< a ¬b ¬c = ≤next (ordtrans a <-osuc ) +x<ny→≤next {x} {y} x<ny | tri≈ ¬a refl ¬c = o≤-refl refl +x<ny→≤next {x} {y} x<ny | tri> ¬a ¬b c = o≤-refl (sym ( x<ny→≡next c x<ny )) + +omax<nomax : {x y : Ordinal} → omax x y o< next (omax x y ) +omax<nomax {x} {y} with trio< x y +omax<nomax {x} {y} | tri< a ¬b ¬c = subst (λ k → osuc y o< k ) nexto≡ (osuc<nx x<nx ) +omax<nomax {x} {y} | tri≈ ¬a refl ¬c = subst (λ k → osuc x o< k ) nexto≡ (osuc<nx x<nx ) +omax<nomax {x} {y} | tri> ¬a ¬b c = subst (λ k → osuc x o< k ) nexto≡ (osuc<nx x<nx ) + +omax<nx : {x y z : Ordinal} → x o< next z → y o< next z → omax x y o< next z +omax<nx {x} {y} {z} x<nz y<nz with trio< x y +omax<nx {x} {y} {z} x<nz y<nz | tri< a ¬b ¬c = osuc<nx y<nz +omax<nx {x} {y} {z} x<nz y<nz | tri≈ ¬a refl ¬c = osuc<nx y<nz +omax<nx {x} {y} {z} x<nz y<nz | tri> ¬a ¬b c = osuc<nx x<nz + +nn<omax : {x nx ny : Ordinal} → x o< next nx → x o< next ny → x o< next (omax nx ny) +nn<omax {x} {nx} {ny} xnx xny with trio< nx ny +nn<omax {x} {nx} {ny} xnx xny | tri< a ¬b ¬c = subst (λ k → x o< k ) nexto≡ xny +nn<omax {x} {nx} {ny} xnx xny | tri≈ ¬a refl ¬c = subst (λ k → x o< k ) nexto≡ xny +nn<omax {x} {nx} {ny} xnx xny | tri> ¬a ¬b c = subst (λ k → x o< k ) nexto≡ xnx + +record OrdinalSubset (maxordinal : Ordinal) : Set (suc n) where + field + os→ : (x : Ordinal) → x o< maxordinal → Ordinal + os← : Ordinal → Ordinal + os←limit : (x : Ordinal) → os← x o< maxordinal + os-iso← : {x : Ordinal} → os→ (os← x) (os←limit x) ≡ x + os-iso→ : {x : Ordinal} → (lt : x o< maxordinal ) → os← (os→ x lt) ≡ x + +module o≤-Reasoning {n : Level} (O : Ordinals {n} ) where + + -- open inOrdinal O + + resp-o< : _o<_ Respects₂ _≡_ + resp-o< = resp₂ _o<_ + + trans1 : {i j k : Ordinal} → i o< j → j o< osuc k → i o< k + trans1 {i} {j} {k} i<j j<ok with osuc-≡< j<ok + trans1 {i} {j} {k} i<j j<ok | case1 refl = i<j + trans1 {i} {j} {k} i<j j<ok | case2 j<k = ordtrans i<j j<k + + trans2 : {i j k : Ordinal} → i o< osuc j → j o< k → i o< k + trans2 {i} {j} {k} i<oj j<k with osuc-≡< i<oj + trans2 {i} {j} {k} i<oj j<k | case1 refl = j<k + trans2 {i} {j} {k} i<oj j<k | case2 i<j = ordtrans i<j j<k + + open import Relation.Binary.Reasoning.Base.Triple + (Preorder.isPreorder OrdPreorder) + ordtrans --<-trans + (resp₂ _o<_) --(resp₂ _<_) + (λ x → ordtrans x <-osuc ) --<⇒≤ + trans1 --<-transˡ + trans2 --<-transʳ + public + -- hiding (_≈⟨_⟩_) +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Ordinals.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,62 @@ +open import Level +module Ordinals where + +open import zf + +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +open import Data.Empty +open import Relation.Binary.PropositionalEquality +open import logic +open import nat +open import Data.Unit using ( ⊤ ) +open import Relation.Nullary +open import Relation.Binary +open import Relation.Binary.Core + +record Oprev {n : Level} (ord : Set n) (osuc : ord → ord ) (x : ord ) : Set (suc n) where + field + oprev : ord + oprev=x : osuc oprev ≡ x + +record IsOrdinals {n : Level} (ord : Set n) (o∅ : ord ) (osuc : ord → ord ) (_o<_ : ord → ord → Set n) (next : ord → ord ) : Set (suc (suc n)) where + field + ordtrans : {x y z : ord } → x o< y → y o< z → x o< z + trio< : Trichotomous {n} _≡_ _o<_ + ¬x<0 : { x : ord } → ¬ ( x o< o∅ ) + <-osuc : { x : ord } → x o< osuc x + osuc-≡< : { a x : ord } → x o< osuc a → (x ≡ a ) ∨ (x o< a) + Oprev-p : ( x : ord ) → Dec ( Oprev ord osuc x ) + TransFinite : { ψ : ord → Set (suc n) } + → ( (x : ord) → ( (y : ord ) → y o< x → ψ y ) → ψ x ) + → ∀ (x : ord) → ψ x + + +record IsNext {n : Level } (ord : Set n) (o∅ : ord ) (osuc : ord → ord ) (_o<_ : ord → ord → Set n) (next : ord → ord ) : Set (suc (suc n)) where + field + x<nx : { y : ord } → (y o< next y ) + osuc<nx : { x y : ord } → x o< next y → osuc x o< next y + ¬nx<nx : {x y : ord} → y o< x → x o< next y → ¬ ((z : ord) → ¬ (x ≡ osuc z)) + +record Ordinals {n : Level} : Set (suc (suc n)) where + field + Ordinal : Set n + o∅ : Ordinal + osuc : Ordinal → Ordinal + _o<_ : Ordinal → Ordinal → Set n + next : Ordinal → Ordinal + isOrdinal : IsOrdinals Ordinal o∅ osuc _o<_ next + isNext : IsNext Ordinal o∅ osuc _o<_ next + +module inOrdinal {n : Level} (O : Ordinals {n} ) where + + open Ordinals O + open IsOrdinals isOrdinal + open IsNext isNext + + TransFinite0 : { ψ : Ordinal → Set n } + → ( (x : Ordinal) → ( (y : Ordinal ) → y o< x → ψ y ) → ψ x ) + → ∀ (x : Ordinal) → ψ x + TransFinite0 {ψ} ind x = lower (TransFinite {λ y → Lift (suc n) ( ψ y)} ind1 x) where + ind1 : (z : Ordinal) → ((y : Ordinal) → y o< z → Lift (suc n) (ψ y)) → Lift (suc n) (ψ z) + ind1 z prev = lift (ind z (λ y y<z → lower (prev y y<z ) )) +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Topology.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,109 @@ +open import Level +open import Ordinals +module Topology {n : Level } (O : Ordinals {n}) where + +open import zf +open import logic +open _∧_ +open _∨_ +open Bool + +import OD +open import Relation.Nullary +open import Data.Empty +open import Relation.Binary.Core +open import Relation.Binary.PropositionalEquality +import BAlgbra +open BAlgbra O +open inOrdinal O +open OD O +open OD.OD +open ODAxiom odAxiom +import OrdUtil +import ODUtil +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +open OrdUtil O +open ODUtil O + +import ODC +open ODC O + +open import filter + +record Toplogy ( L : HOD ) : Set (suc n) where + field + OS : HOD + OS⊆PL : OS ⊆ Power L + o∪ : { P : HOD } → P ⊆ OS → OS ∋ Union P + o∩ : { p q : HOD } → OS ∋ p → OS ∋ q → OS ∋ (p ∩ q) + +open Toplogy + +record _covers_ ( P q : HOD ) : Set (suc n) where + field + cover : {x : HOD} → q ∋ x → HOD + P∋cover : {x : HOD} → {lt : q ∋ x} → P ∋ cover lt + isCover : {x : HOD} → {lt : q ∋ x} → cover lt ∋ x + +-- Base +-- The elements of B cover X ; For any U , V ∈ B and any point x ∈ U ∩ V there is a W ∈ B such that +-- W ⊆ U ∩ V and x ∈ W . + +data genTop (P : HOD) : HOD → Set (suc n) where + gi : {x : HOD} → P ∋ x → genTop P x + g∩ : {x y : HOD} → genTop P x → genTop P y → genTop P (x ∩ y) + g∪ : {Q x : HOD} → Q ⊆ P → genTop P (Union Q) + +-- Limit point + +record LP ( L S x : HOD ) (top : Toplogy L) (S⊆PL : S ⊆ Power L ) ( S∋x : S ∋ x ) : Set (suc n) where + field + neip : {y : HOD} → OS top ∋ y → y ∋ x → HOD + isNeip : {y : HOD} → (o∋y : OS top ∋ y ) → (y∋x : y ∋ x ) → ¬ ( x ≡ neip o∋y y∋x) ∧ ( y ∋ neip o∋y y∋x ) + +-- Finite Intersection Property + +data Finite-∩ (S : HOD) : HOD → Set (suc n) where + fin-∩e : {x : HOD} → S ∋ x → Finite-∩ S x + fin-∩ : {x y : HOD} → Finite-∩ S x → Finite-∩ S y → Finite-∩ S (x ∩ y) + +record FIP ( L P : HOD ) : Set (suc n) where + field + fipS⊆PL : P ⊆ Power L + fip≠φ : { x : HOD } → Finite-∩ P x → ¬ ( x ≡ od∅ ) + +-- Compact + +data Finite-∪ (S : HOD) : HOD → Set (suc n) where + fin-∪e : {x : HOD} → S ∋ x → Finite-∪ S x + fin-∪ : {x y : HOD} → Finite-∪ S x → Finite-∪ S y → Finite-∪ S (x ∪ y) + +record Compact ( L P : HOD ) : Set (suc n) where + field + finCover : {X y : HOD} → X covers P → P ∋ y → HOD + isFinCover : {X y : HOD} → (xp : X covers P ) → (P∋y : P ∋ y ) → finCover xp P∋y ∋ y + isFininiteCover : {X y : HOD} → (xp : X covers P ) → (P∋y : P ∋ y ) → Finite-∪ X (finCover xp P∋y ) + +-- FIP is Compact + +FIP→Compact : {L P : HOD} → Tolopogy L → FIP L P → Compact L P +FIP→Compact = ? + +Compact→FIP : {L P : HOD} → Tolopogy L → Compact L P → FIP L P +Compact→FIP = ? + +-- Product Topology + +_Top⊗_ : {P Q : HOD} → Topology P → Tolopogy Q → Topology ( P ⊗ Q ) +_Top⊗_ = ? + +-- existence of Ultra Filter + +-- Ultra Filter has limit point + +-- FIP is UFL + +-- Product of UFL has limit point (Tychonoff) +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/VL.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,86 @@ +open import Level +open import Ordinals +module VL {n : Level } (O : Ordinals {n}) where + +open import zf +open import logic +import OD +open import Relation.Nullary +open import Relation.Binary +open import Data.Empty +open import Relation.Binary +open import Relation.Binary.Core +open import Relation.Binary.PropositionalEquality +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +import BAlgbra +open BAlgbra O +open inOrdinal O +import OrdUtil +import ODUtil +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +open OrdUtil O +open ODUtil O + +open OD O +open OD.OD +open ODAxiom odAxiom +-- import ODC +open _∧_ +open _∨_ +open Bool +open HOD + +-- The cumulative hierarchy +-- V 0 := ∅ +-- V α + 1 := P ( V α ) +-- V α := ⋃ { V β | β < α } + +V : ( β : Ordinal ) → HOD +V β = TransFinite V1 β where + V1 : (x : Ordinal ) → ( ( y : Ordinal) → y o< x → HOD ) → HOD + V1 x V0 with trio< x o∅ + V1 x V0 | tri< a ¬b ¬c = ⊥-elim ( ¬x<0 a) + V1 x V0 | tri≈ ¬a refl ¬c = Ord o∅ + V1 x V0 | tri> ¬a ¬b c with Oprev-p x + V1 x V0 | tri> ¬a ¬b c | yes p = Power ( V0 (Oprev.oprev p ) (subst (λ k → _ o< k) (Oprev.oprev=x p) <-osuc )) + V1 x V0 | tri> ¬a ¬b c | no ¬p = + record { od = record { def = λ y → (y o< x ) ∧ ((lt : y o< x ) → odef (V0 y lt) x ) } ; odmax = x; <odmax = λ {x} lt → proj1 lt } + +-- +-- L ⊆ HOD ⊆ V +-- +-- HOD=V : (x : HOD) → V (odmax x) ∋ x +-- HOD=V x = {!!} + +-- +-- Definition for Power Set +-- +record Definitions : Set (suc n) where + field + Definition : HOD → HOD + +open Definitions + +Df : Definitions → (x : HOD) → HOD +Df D x = Power x ∩ Definition D x + +-- The constructible Sets +-- L 0 := ∅ +-- L α + 1 := Df (L α) -- Definable Power Set +-- V α := ⋃ { L β | β < α } + +L : ( β : Ordinal ) → Definitions → HOD +L β D = TransFinite L1 β where + L1 : (x : Ordinal ) → ( ( y : Ordinal) → y o< x → HOD ) → HOD + L1 x L0 with trio< x o∅ + L1 x L0 | tri< a ¬b ¬c = ⊥-elim ( ¬x<0 a) + L1 x L0 | tri≈ ¬a refl ¬c = Ord o∅ + L1 x L0 | tri> ¬a ¬b c with Oprev-p x + L1 x L0 | tri> ¬a ¬b c | yes p = Df D ( L0 (Oprev.oprev p ) (subst (λ k → _ o< k) (Oprev.oprev=x p) <-osuc )) + L1 x L0 | tri> ¬a ¬b c | no ¬p = + record { od = record { def = λ y → (y o< x ) ∧ ((lt : y o< x ) → odef (L0 y lt) x ) } ; odmax = x; <odmax = λ {x} lt → proj1 lt } + +V=L0 : Set (suc n) +V=L0 = (x : Ordinal) → V x ≡ L x record { Definition = λ y → y }
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/cardinal.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,95 @@ +open import Level +open import Ordinals +module cardinal {n : Level } (O : Ordinals {n}) where + +open import zf +open import logic +import OD +import ODC +import OPair +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +open import Relation.Binary.PropositionalEquality +open import Data.Nat.Properties +open import Data.Empty +open import Relation.Nullary +open import Relation.Binary +open import Relation.Binary.Core + +open inOrdinal O +open OD O +open OD.OD +open OPair O +open ODAxiom odAxiom + +import OrdUtil +import ODUtil +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +open OrdUtil O +open ODUtil O + + +open _∧_ +open _∨_ +open Bool +open _==_ + +open HOD + +-- _⊗_ : (A B : HOD) → HOD +-- A ⊗ B = Union ( Replace B (λ b → Replace A (λ a → < a , b > ) )) + +Func : OD +Func = record { def = λ x → def ZFProduct x + ∧ ( (a b c : Ordinal) → odef (* x) (& < * a , * b >) ∧ odef (* x) (& < * a , * c >) → b ≡ c ) } + +FuncP : ( A B : HOD ) → HOD +FuncP A B = record { od = record { def = λ x → odef (ZFP A B) x + ∧ ( (x : Ordinal ) (p q : odef (ZFP A B ) x ) → pi1 (proj1 p) ≡ pi1 (proj1 q) → pi2 (proj1 p) ≡ pi2 (proj1 q) ) } + ; odmax = odmax (ZFP A B) ; <odmax = λ lt → <odmax (ZFP A B) (proj1 lt) } + +record Injection (A B : Ordinal ) : Set n where + field + i→ : (x : Ordinal ) → odef (* A) x → Ordinal + iB : (x : Ordinal ) → ( lt : odef (* A) x ) → odef (* B) ( i→ x lt ) + iiso : (x y : Ordinal ) → ( ltx : odef (* A) x ) ( lty : odef (* A) y ) → i→ x ltx ≡ i→ y lty → x ≡ y + +record Bijection (A B : Ordinal ) : Set n where + field + fun← : (x : Ordinal ) → odef (* A) x → Ordinal + fun→ : (x : Ordinal ) → odef (* B) x → Ordinal + funB : (x : Ordinal ) → ( lt : odef (* A) x ) → odef (* B) ( fun← x lt ) + funA : (x : Ordinal ) → ( lt : odef (* B) x ) → odef (* A) ( fun→ x lt ) + fiso← : (x : Ordinal ) → ( lt : odef (* B) x ) → fun← ( fun→ x lt ) ( funA x lt ) ≡ x + fiso→ : (x : Ordinal ) → ( lt : odef (* A) x ) → fun→ ( fun← x lt ) ( funB x lt ) ≡ x + +Bernstein : {a b : Ordinal } → Injection a b → Injection b a → Bijection a b +Bernstein = {!!} + +_=c=_ : ( A B : HOD ) → Set n +A =c= B = Bijection ( & A ) ( & B ) + +_c<_ : ( A B : HOD ) → Set n +A c< B = ¬ ( Injection (& A) (& B) ) + +Card : OD +Card = record { def = λ x → (a : Ordinal) → a o< x → ¬ Bijection a x } + +record Cardinal (a : Ordinal ) : Set (suc n) where + field + card : Ordinal + ciso : Bijection a card + cmax : (x : Ordinal) → card o< x → ¬ Bijection a x + +Cardinal∈ : { s : HOD } → { t : Ordinal } → Ord t ∋ s → s c< Ord t +Cardinal∈ = {!!} + +Cardinal⊆ : { s t : HOD } → s ⊆ t → ( s c< t ) ∨ ( s =c= t ) +Cardinal⊆ = {!!} + +Cantor1 : { u : HOD } → u c< Power u +Cantor1 = {!!} + +Cantor2 : { u : HOD } → ¬ ( u =c= Power u ) +Cantor2 = {!!}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/filter.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,195 @@ +open import Level +open import Ordinals +module filter {n : Level } (O : Ordinals {n}) where + +open import zf +open import logic +import OD + +open import Relation.Nullary +open import Data.Empty +open import Relation.Binary.Core +open import Relation.Binary.PropositionalEquality +import BAlgbra + +open BAlgbra O + +open inOrdinal O +open OD O +open OD.OD +open ODAxiom odAxiom + +import OrdUtil +import ODUtil +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +open OrdUtil O +open ODUtil O + + +import ODC +open ODC O + +open _∧_ +open _∨_ +open Bool + +-- Kunen p.76 and p.53, we use ⊆ +record Filter ( L : HOD ) : Set (suc n) where + field + filter : HOD + f⊆PL : filter ⊆ Power L + filter1 : { p q : HOD } → q ⊆ L → filter ∋ p → p ⊆ q → filter ∋ q + filter2 : { p q : HOD } → filter ∋ p → filter ∋ q → filter ∋ (p ∩ q) + +open Filter + +record prime-filter { L : HOD } (P : Filter L) : Set (suc (suc n)) where + field + proper : ¬ (filter P ∋ od∅) + prime : {p q : HOD } → filter P ∋ (p ∪ q) → ( filter P ∋ p ) ∨ ( filter P ∋ q ) + +record ultra-filter { L : HOD } (P : Filter L) : Set (suc (suc n)) where + field + proper : ¬ (filter P ∋ od∅) + ultra : {p : HOD } → p ⊆ L → ( filter P ∋ p ) ∨ ( filter P ∋ ( L \ p) ) + +open _⊆_ + +∈-filter : {L p : HOD} → (P : Filter L ) → filter P ∋ p → p ⊆ L +∈-filter {L} {p} P lt = power→⊆ L p ( incl (f⊆PL P) lt ) + +∪-lemma1 : {L p q : HOD } → (p ∪ q) ⊆ L → p ⊆ L +∪-lemma1 {L} {p} {q} lt = record { incl = λ {x} p∋x → incl lt (case1 p∋x) } + +∪-lemma2 : {L p q : HOD } → (p ∪ q) ⊆ L → q ⊆ L +∪-lemma2 {L} {p} {q} lt = record { incl = λ {x} p∋x → incl lt (case2 p∋x) } + +q∩q⊆q : {p q : HOD } → (q ∩ p) ⊆ q +q∩q⊆q = record { incl = λ lt → proj1 lt } + +open HOD + +----- +-- +-- ultra filter is prime +-- + +filter-lemma1 : {L : HOD} → (P : Filter L) → ∀ {p q : HOD } → ultra-filter P → prime-filter P +filter-lemma1 {L} P u = record { + proper = ultra-filter.proper u + ; prime = lemma3 + } where + lemma3 : {p q : HOD} → filter P ∋ (p ∪ q) → ( filter P ∋ p ) ∨ ( filter P ∋ q ) + lemma3 {p} {q} lt with ultra-filter.ultra u (∪-lemma1 (∈-filter P lt) ) + ... | case1 p∈P = case1 p∈P + ... | case2 ¬p∈P = case2 (filter1 P {q ∩ (L \ p)} (∪-lemma2 (∈-filter P lt)) lemma7 lemma8) where + lemma5 : ((p ∪ q ) ∩ (L \ p)) =h= (q ∩ (L \ p)) + lemma5 = record { eq→ = λ {x} lt → ⟪ lemma4 x lt , proj2 lt ⟫ + ; eq← = λ {x} lt → ⟪ case2 (proj1 lt) , proj2 lt ⟫ + } where + lemma4 : (x : Ordinal ) → odef ((p ∪ q) ∩ (L \ p)) x → odef q x + lemma4 x lt with proj1 lt + lemma4 x lt | case1 px = ⊥-elim ( proj2 (proj2 lt) px ) + lemma4 x lt | case2 qx = qx + lemma6 : filter P ∋ ((p ∪ q ) ∩ (L \ p)) + lemma6 = filter2 P lt ¬p∈P + lemma7 : filter P ∋ (q ∩ (L \ p)) + lemma7 = subst (λ k → filter P ∋ k ) (==→o≡ lemma5 ) lemma6 + lemma8 : (q ∩ (L \ p)) ⊆ q + lemma8 = q∩q⊆q + +----- +-- +-- if Filter contains L, prime filter is ultra +-- + +filter-lemma2 : {L : HOD} → (P : Filter L) → filter P ∋ L → prime-filter P → ultra-filter P +filter-lemma2 {L} P f∋L prime = record { + proper = prime-filter.proper prime + ; ultra = λ {p} p⊆L → prime-filter.prime prime (lemma p p⊆L) + } where + open _==_ + p+1-p=1 : {p : HOD} → p ⊆ L → L =h= (p ∪ (L \ p)) + eq→ (p+1-p=1 {p} p⊆L) {x} lt with ODC.decp O (odef p x) + eq→ (p+1-p=1 {p} p⊆L) {x} lt | yes p∋x = case1 p∋x + eq→ (p+1-p=1 {p} p⊆L) {x} lt | no ¬p = case2 ⟪ lt , ¬p ⟫ + eq← (p+1-p=1 {p} p⊆L) {x} ( case1 p∋x ) = subst (λ k → odef L k ) &iso (incl p⊆L ( subst (λ k → odef p k) (sym &iso) p∋x )) + eq← (p+1-p=1 {p} p⊆L) {x} ( case2 ¬p ) = proj1 ¬p + lemma : (p : HOD) → p ⊆ L → filter P ∋ (p ∪ (L \ p)) + lemma p p⊆L = subst (λ k → filter P ∋ k ) (==→o≡ (p+1-p=1 p⊆L)) f∋L + +record Dense (P : HOD ) : Set (suc n) where + field + dense : HOD + d⊆P : dense ⊆ Power P + dense-f : HOD → HOD + dense-d : { p : HOD} → p ⊆ P → dense ∋ dense-f p + dense-p : { p : HOD} → p ⊆ P → p ⊆ (dense-f p) + +record Ideal ( L : HOD ) : Set (suc n) where + field + ideal : HOD + i⊆PL : ideal ⊆ Power L + ideal1 : { p q : HOD } → q ⊆ L → ideal ∋ p → q ⊆ p → ideal ∋ q + ideal2 : { p q : HOD } → ideal ∋ p → ideal ∋ q → ideal ∋ (p ∪ q) + +open Ideal + +proper-ideal : {L : HOD} → (P : Ideal L ) → {p : HOD} → Set n +proper-ideal {L} P {p} = ideal P ∋ od∅ + +prime-ideal : {L : HOD} → Ideal L → ∀ {p q : HOD } → Set n +prime-ideal {L} P {p} {q} = ideal P ∋ ( p ∩ q) → ( ideal P ∋ p ) ∨ ( ideal P ∋ q ) + +---- +-- +-- Filter/Ideal without ZF +-- L have to be a Latice +-- + +record F-Filter {n : Level} (L : Set n) (PL : (L → Set n) → Set n) ( _⊆_ : L → L → Set n) (_∩_ : L → L → L ) : Set (suc n) where + field + filter : L → Set n + f⊆P : PL filter + filter1 : { p q : L } → PL (λ x → q ⊆ x ) → filter p → p ⊆ q → filter q + filter2 : { p q : L } → filter p → filter q → filter (p ∩ q) + +Filter-is-F : {L : HOD} → (f : Filter L ) → F-Filter HOD (λ p → (x : HOD) → p x → x ⊆ L ) _⊆_ _∩_ +Filter-is-F {L} f = record { + filter = λ x → Lift (suc n) ((filter f) ∋ x) + ; f⊆P = λ x f∋x → power→⊆ _ _ (incl ( f⊆PL f ) (lower f∋x )) + ; filter1 = λ {p} {q} q⊆L f∋p p⊆q → lift ( filter1 f (q⊆L q refl-⊆) (lower f∋p) p⊆q) + ; filter2 = λ {p} {q} f∋p f∋q → lift ( filter2 f (lower f∋p) (lower f∋q)) + } + +record F-Dense {n : Level} (L : Set n) (PL : (L → Set n) → Set n) ( _⊆_ : L → L → Set n) (_∩_ : L → L → L ) : Set (suc n) where + field + dense : L → Set n + d⊆P : PL dense + dense-f : L → L + dense-d : { p : L} → PL (λ x → p ⊆ x ) → dense ( dense-f p ) + dense-p : { p : L} → PL (λ x → p ⊆ x ) → p ⊆ (dense-f p) + +Dense-is-F : {L : HOD} → (f : Dense L ) → F-Dense HOD (λ p → (x : HOD) → p x → x ⊆ L ) _⊆_ _∩_ +Dense-is-F {L} f = record { + dense = λ x → Lift (suc n) ((dense f) ∋ x) + ; d⊆P = λ x f∋x → power→⊆ _ _ (incl ( d⊆P f ) (lower f∋x )) + ; dense-f = λ x → dense-f f x + ; dense-d = λ {p} d → lift ( dense-d f (d p refl-⊆ ) ) + ; dense-p = λ {p} d → dense-p f (d p refl-⊆) + } where open Dense + + +record GenericFilter (P : HOD) : Set (suc n) where + field + genf : Filter P + generic : (D : Dense P ) → ¬ ( (Dense.dense D ∩ Filter.filter genf ) ≡ od∅ ) + +record F-GenericFilter {n : Level} (L : Set n) (PL : (L → Set n) → Set n) ( _⊆_ : L → L → Set n) (_∩_ : L → L → L ) : Set (suc n) where + field + GFilter : F-Filter L PL _⊆_ _∩_ + Intersection : (D : F-Dense L PL _⊆_ _∩_ ) → { x : L } → F-Dense.dense D x → L + Generic : (D : F-Dense L PL _⊆_ _∩_ ) → { x : L } → ( y : F-Dense.dense D x) → F-Filter.filter GFilter (Intersection D y ) +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/generic-filter.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,263 @@ +open import Level +open import Ordinals +module generic-filter {n : Level } (O : Ordinals {n}) where + +import filter +open import zf +open import logic +-- open import partfunc {n} O +import OD + +open import Relation.Nullary +open import Relation.Binary +open import Data.Empty +open import Relation.Binary +open import Relation.Binary.Core +open import Relation.Binary.PropositionalEquality +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +import BAlgbra + +open BAlgbra O + +open inOrdinal O +open OD O +open OD.OD +open ODAxiom odAxiom +import OrdUtil +import ODUtil +open Ordinals.Ordinals O +open Ordinals.IsOrdinals isOrdinal +open Ordinals.IsNext isNext +open OrdUtil O +open ODUtil O + + +import ODC + +open filter O + +open _∧_ +open _∨_ +open Bool + + +open HOD + +------- +-- the set of finite partial functions from ω to 2 +-- +-- + +open import Data.List hiding (filter) +open import Data.Maybe + +import OPair +open OPair O + +ODSuc : (y : HOD) → infinite ∋ y → HOD +ODSuc y lt = Union (y , (y , y)) + +data Hω2 : (i : Nat) ( x : Ordinal ) → Set n where + hφ : Hω2 0 o∅ + h0 : {i : Nat} {x : Ordinal } → Hω2 i x → + Hω2 (Suc i) (& (Union ((< nat→ω i , nat→ω 0 >) , * x ))) + h1 : {i : Nat} {x : Ordinal } → Hω2 i x → + Hω2 (Suc i) (& (Union ((< nat→ω i , nat→ω 1 >) , * x ))) + he : {i : Nat} {x : Ordinal } → Hω2 i x → + Hω2 (Suc i) x + +record Hω2r (x : Ordinal) : Set n where + field + count : Nat + hω2 : Hω2 count x + +open Hω2r + +HODω2 : HOD +HODω2 = record { od = record { def = λ x → Hω2r x } ; odmax = next o∅ ; <odmax = odmax0 } where + P : (i j : Nat) (x : Ordinal ) → HOD + P i j x = ((nat→ω i , nat→ω i) , (nat→ω i , nat→ω j)) , * x + nat1 : (i : Nat) (x : Ordinal) → & (nat→ω i) o< next x + nat1 i x = next< nexto∅ ( <odmax infinite (ω∋nat→ω {i})) + lemma1 : (i j : Nat) (x : Ordinal ) → osuc (& (P i j x)) o< next x + lemma1 i j x = osuc<nx (pair-<xy (pair-<xy (pair-<xy (nat1 i x) (nat1 i x) ) (pair-<xy (nat1 i x) (nat1 j x) ) ) + (subst (λ k → k o< next x) (sym &iso) x<nx )) + lemma : (i j : Nat) (x : Ordinal ) → & (Union (P i j x)) o< next x + lemma i j x = next< (lemma1 i j x ) ho< + odmax0 : {y : Ordinal} → Hω2r y → y o< next o∅ + odmax0 {y} r with hω2 r + ... | hφ = x<nx + ... | h0 {i} {x} t = next< (odmax0 record { count = i ; hω2 = t }) (lemma i 0 x) + ... | h1 {i} {x} t = next< (odmax0 record { count = i ; hω2 = t }) (lemma i 1 x) + ... | he {i} {x} t = next< (odmax0 record { count = i ; hω2 = t }) x<nx + +3→Hω2 : List (Maybe Two) → HOD +3→Hω2 t = list→hod t 0 where + list→hod : List (Maybe Two) → Nat → HOD + list→hod [] _ = od∅ + list→hod (just i0 ∷ t) i = Union (< nat→ω i , nat→ω 0 > , ( list→hod t (Suc i) )) + list→hod (just i1 ∷ t) i = Union (< nat→ω i , nat→ω 1 > , ( list→hod t (Suc i) )) + list→hod (nothing ∷ t) i = list→hod t (Suc i ) + +Hω2→3 : (x : HOD) → HODω2 ∋ x → List (Maybe Two) +Hω2→3 x = lemma where + lemma : { y : Ordinal } → Hω2r y → List (Maybe Two) + lemma record { count = 0 ; hω2 = hφ } = [] + lemma record { count = (Suc i) ; hω2 = (h0 hω3) } = just i0 ∷ lemma record { count = i ; hω2 = hω3 } + lemma record { count = (Suc i) ; hω2 = (h1 hω3) } = just i1 ∷ lemma record { count = i ; hω2 = hω3 } + lemma record { count = (Suc i) ; hω2 = (he hω3) } = nothing ∷ lemma record { count = i ; hω2 = hω3 } + +ω→2 : HOD +ω→2 = Power infinite + +ω→2f : (x : HOD) → ω→2 ∋ x → Nat → Two +ω→2f x lt n with ODC.∋-p O x (nat→ω n) +ω→2f x lt n | yes p = i1 +ω→2f x lt n | no ¬p = i0 + +fω→2-sel : ( f : Nat → Two ) (x : HOD) → Set n +fω→2-sel f x = (infinite ∋ x) ∧ ( (lt : odef infinite (& x) ) → f (ω→nat x lt) ≡ i1 ) + +fω→2 : (Nat → Two) → HOD +fω→2 f = Select infinite (fω→2-sel f) + +open _==_ + +import Axiom.Extensionality.Propositional +postulate f-extensionality : { n m : Level} → Axiom.Extensionality.Propositional.Extensionality n m + +ω2∋f : (f : Nat → Two) → ω→2 ∋ fω→2 f +ω2∋f f = power← infinite (fω→2 f) (λ {x} lt → proj1 ((proj2 (selection {fω→2-sel f} {infinite} )) lt)) + +ω→2f≡i1 : (X i : HOD) → (iω : infinite ∋ i) → (lt : ω→2 ∋ X ) → ω→2f X lt (ω→nat i iω) ≡ i1 → X ∋ i +ω→2f≡i1 X i iω lt eq with ODC.∋-p O X (nat→ω (ω→nat i iω)) +ω→2f≡i1 X i iω lt eq | yes p = subst (λ k → X ∋ k ) (nat→ω-iso iω) p + +open _⊆_ + +-- someday ... +-- postulate +-- ω→2f-iso : (X : HOD) → ( lt : ω→2 ∋ X ) → fω→2 ( ω→2f X lt ) =h= X +-- fω→2-iso : (f : Nat → Two) → ω→2f ( fω→2 f ) (ω2∋f f) ≡ f + +record CountableOrdinal : Set (suc (suc n)) where + field + ctl→ : Nat → Ordinal + ctl← : Ordinal → Nat + ctl-iso→ : { x : Ordinal } → ctl→ (ctl← x ) ≡ x + ctl-iso← : { x : Nat } → ctl← (ctl→ x ) ≡ x + +record CountableHOD : Set (suc (suc n)) where + field + mhod : HOD + mtl→ : Nat → Ordinal + mtl→∈P : (i : Nat) → odef mhod (mtl→ i) + mtl← : (x : Ordinal) → odef mhod x → Nat + mtl-iso→ : { x : Ordinal } → (lt : odef mhod x ) → mtl→ (mtl← x lt ) ≡ x + mtl-iso← : { x : Nat } → mtl← (mtl→ x ) (mtl→∈P x) ≡ x + + +open CountableOrdinal +open CountableHOD + +---- +-- a(n) ∈ M +-- ∃ q ∈ Power P → q ∈ a(n) ∧ p(n) ⊆ q +-- +PGHOD : (i : Nat) → (C : CountableOrdinal) → (P : HOD) → (p : Ordinal) → HOD +PGHOD i C P p = record { od = record { def = λ x → + odef (Power P) x ∧ odef (* (ctl→ C i)) x ∧ ( (y : Ordinal ) → odef (* p) y → odef (* x) y ) } + ; odmax = odmax (Power P) ; <odmax = λ {y} lt → <odmax (Power P) (proj1 lt) } + +--- +-- p(n+1) = if PGHOD n qn otherwise p(n) +-- +next-p : (C : CountableOrdinal) (P : HOD ) (i : Nat) → (p : Ordinal) → Ordinal +next-p C P i p with is-o∅ ( & (PGHOD i C P p)) +next-p C P i p | yes y = p +next-p C P i p | no not = & (ODC.minimal O (PGHOD i C P p ) (λ eq → not (=od∅→≡o∅ eq))) -- axiom of choice + +--- +-- ascendant search on p(n) +-- +find-p : (C : CountableOrdinal) (P : HOD ) (i : Nat) → (x : Ordinal) → Ordinal +find-p C P Zero x = x +find-p C P (Suc i) x = find-p C P i ( next-p C P i x ) + +--- +-- G = { r ∈ Power P | ∃ n → r ⊆ p(n) } +-- +record PDN (C : CountableOrdinal) (P : HOD ) (x : Ordinal) : Set n where + field + gr : Nat + pn<gr : (y : Ordinal) → odef (* x) y → odef (* (find-p C P gr o∅)) y + x∈PP : odef (Power P) x + +open PDN + +--- +-- G as a HOD +-- +PDHOD : (C : CountableOrdinal) → (P : HOD ) → HOD +PDHOD C P = record { od = record { def = λ x → PDN C P x } + ; odmax = odmax (Power P) ; <odmax = λ {y} lt → <odmax (Power P) {y} (PDN.x∈PP lt) } + +open PDN + +---- +-- Generic Filter on Power P for HOD's Countable Ordinal (G ⊆ Power P ≡ G i.e. Nat → P → Set ) +-- +-- p 0 ≡ ∅ +-- p (suc n) = if ∃ q ∈ * ( ctl→ n ) ∧ p n ⊆ q → q (axiom of choice) +--- else p n + +P∅ : {P : HOD} → odef (Power P) o∅ +P∅ {P} = subst (λ k → odef (Power P) k ) ord-od∅ (lemma o∅ o∅≡od∅) where + lemma : (x : Ordinal ) → * x ≡ od∅ → odef (Power P) (& od∅) + lemma x eq = power← P od∅ (λ {x} lt → ⊥-elim (¬x<0 lt )) +x<y→∋ : {x y : Ordinal} → odef (* x) y → * x ∋ * y +x<y→∋ {x} {y} lt = subst (λ k → odef (* x) k ) (sym &iso) lt + +P-GenericFilter : (C : CountableOrdinal) → (P : HOD ) → GenericFilter P +P-GenericFilter C P = record { + genf = record { filter = PDHOD C P ; f⊆PL = f⊆PL ; filter1 = {!!} ; filter2 = {!!} } + ; generic = λ D → {!!} + } where + find-p-⊆P : (C : CountableOrdinal) (P : HOD ) (i : Nat) → (x y : Ordinal) → odef (Power P) x → odef (* (find-p C P i x)) y → odef P y + find-p-⊆P C P Zero x y Px Py = subst (λ k → odef P k ) &iso + ( incl (ODC.power→⊆ O P (* x) (d→∋ (Power P) Px)) (x<y→∋ Py)) + find-p-⊆P C P (Suc i) x y Px Py = find-p-⊆P C P i (next-p C P i x) y {!!} {!!} + f⊆PL : PDHOD C P ⊆ Power P + f⊆PL = record { incl = λ {x} lt → power← P x (λ {y} y<x → + find-p-⊆P C P (gr lt) o∅ (& y) P∅ (pn<gr lt (& y) (subst (λ k → odef k (& y)) (sym *iso) y<x))) } + + +open GenericFilter +open Filter + +record Incompatible (P : HOD ) : Set (suc (suc n)) where + field + except : HOD → ( HOD ∧ HOD ) + incompatible : { p : HOD } → Power P ∋ p → Power P ∋ proj1 (except p ) → Power P ∋ proj2 (except p ) + → ( p ⊆ proj1 (except p) ) ∧ ( p ⊆ proj2 (except p) ) + → ∀ ( r : HOD ) → Power P ∋ r → ¬ (( proj1 (except p) ⊆ r ) ∧ ( proj2 (except p) ⊆ r )) + +lemma725 : (M : CountableHOD ) (C : CountableOrdinal) (P : HOD ) → mhod M ∋ Power P + → Incompatible P → ¬ ( mhod M ∋ filter ( genf ( P-GenericFilter C P ))) +lemma725 = {!!} + +lemma725-1 : Incompatible HODω2 +lemma725-1 = {!!} + +lemma726 : (C : CountableOrdinal) (P : HOD ) + → Union ( filter ( genf ( P-GenericFilter C HODω2 ))) =h= ω→2 +lemma726 = {!!} + +-- +-- val x G = { val y G | ∃ p → G ∋ p → x ∋ < y , p > } +-- +-- W (ω , H ( ω , 2 )) = { p ∈ ( Nat → H (ω , 2) ) | { i ∈ Nat → p i ≠ i1 } is finite } +-- + + +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/logic.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,57 @@ +module logic where + +open import Level +open import Relation.Nullary +open import Relation.Binary hiding (_⇔_ ) +open import Data.Empty + +data One {n : Level } : Set n where + OneObj : One + +data Two : Set where + i0 : Two + i1 : Two + +data Bool : Set where + true : Bool + false : Bool + +record _∧_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where + constructor ⟪_,_⟫ + field + proj1 : A + proj2 : B + +data _∨_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where + case1 : A → A ∨ B + case2 : B → A ∨ B + +_⇔_ : {n m : Level } → ( A : Set n ) ( B : Set m ) → Set (n ⊔ m) +_⇔_ A B = ( A → B ) ∧ ( B → A ) + +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 ) + +double-neg : {n : Level } {A : Set n} → A → ¬ ¬ A +double-neg A notnot = notnot A + +double-neg2 : {n : Level } {A : Set n} → ¬ ¬ ¬ A → ¬ A +double-neg2 notnot A = notnot ( double-neg A ) + +de-morgan : {n : Level } {A B : Set n} → A ∧ B → ¬ ( (¬ A ) ∨ (¬ B ) ) +de-morgan {n} {A} {B} and (case1 ¬A) = ⊥-elim ( ¬A ( _∧_.proj1 and )) +de-morgan {n} {A} {B} and (case2 ¬B) = ⊥-elim ( ¬B ( _∧_.proj2 and )) + +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 + +dont-orb : {n m : Level} {A : Set n} { B : Set m } → A ∨ B → ¬ B → A +dont-orb {A} {B} (case2 b) ¬B = ⊥-elim ( ¬B b ) +dont-orb {A} {B} (case1 a) ¬B = a + + +infixr 130 _∧_ +infixr 140 _∨_ +infixr 150 _⇔_ +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/nat.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,46 @@ +module nat where + +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +open import Data.Empty +open import Relation.Nullary +open import Relation.Binary.PropositionalEquality +open import logic + + +nat-<> : { x y : Nat } → x < y → y < x → ⊥ +nat-<> (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x + +nat-<≡ : { x : Nat } → x < x → ⊥ +nat-<≡ (s≤s lt) = nat-<≡ lt + +nat-≡< : { x y : Nat } → x ≡ y → x < y → ⊥ +nat-≡< refl lt = nat-<≡ lt + +¬a≤a : {la : Nat} → Suc la ≤ la → ⊥ +¬a≤a (s≤s lt) = ¬a≤a lt + +a<sa : {la : Nat} → la < Suc la +a<sa {Zero} = s≤s z≤n +a<sa {Suc la} = s≤s a<sa + +=→¬< : {x : Nat } → ¬ ( x < x ) +=→¬< {Zero} () +=→¬< {Suc x} (s≤s lt) = =→¬< lt + +>→¬< : {x y : Nat } → (x < y ) → ¬ ( y < x ) +>→¬< (s≤s x<y) (s≤s y<x) = >→¬< x<y y<x + +<-∨ : { x y : Nat } → x < Suc y → ( (x ≡ y ) ∨ (x < y) ) +<-∨ {Zero} {Zero} (s≤s z≤n) = case1 refl +<-∨ {Zero} {Suc y} (s≤s lt) = case2 (s≤s z≤n) +<-∨ {Suc x} {Zero} (s≤s ()) +<-∨ {Suc x} {Suc y} (s≤s lt) with <-∨ {x} {y} lt +<-∨ {Suc x} {Suc y} (s≤s lt) | case1 eq = case1 (cong (λ k → Suc k ) eq) +<-∨ {Suc x} {Suc y} (s≤s lt) | case2 lt1 = case2 (s≤s lt1) + +max : (x y : Nat) → Nat +max Zero Zero = Zero +max Zero (Suc x) = (Suc x) +max (Suc x) Zero = (Suc x) +max (Suc x) (Suc y) = Suc ( max x y ) +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ordinal.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,282 @@ +open import Level +module ordinal where + +open import logic +open import nat +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +open import Data.Empty +open import Relation.Binary.PropositionalEquality +open import Relation.Binary.Definitions +open import Data.Nat.Properties +open import Relation.Nullary +open import Relation.Binary.Core + +---- +-- +-- Countable Ordinals +-- + +data OrdinalD {n : Level} : (lv : Nat) → Set n where + Φ : (lv : Nat) → OrdinalD lv + OSuc : (lv : Nat) → OrdinalD {n} lv → OrdinalD lv + +record Ordinal {n : Level} : Set n where + constructor ordinal + field + lv : Nat + ord : OrdinalD {n} lv + +data _d<_ {n : Level} : {lx ly : Nat} → OrdinalD {n} lx → OrdinalD {n} ly → Set n where + Φ< : {lx : Nat} → {x : OrdinalD {n} lx} → Φ lx d< OSuc lx x + s< : {lx : Nat} → {x y : OrdinalD {n} lx} → x d< y → OSuc lx x d< OSuc lx y + +open Ordinal + +_o<_ : {n : Level} ( x y : Ordinal ) → Set n +_o<_ x y = (lv x < lv y ) ∨ ( ord x d< ord y ) + +s<refl : {n : Level } {lx : Nat } { x : OrdinalD {n} lx } → x d< OSuc lx x +s<refl {n} {lv} {Φ lv} = Φ< +s<refl {n} {lv} {OSuc lv x} = s< s<refl + +trio<> : {n : Level} → {lx : Nat} {x : OrdinalD {n} lx } { y : OrdinalD lx } → y d< x → x d< y → ⊥ +trio<> {n} {lx} {.(OSuc lx _)} {.(OSuc lx _)} (s< s) (s< t) = trio<> s t +trio<> {n} {lx} {.(OSuc lx _)} {.(Φ lx)} Φ< () + +d<→lv : {n : Level} {x y : Ordinal {n}} → ord x d< ord y → lv x ≡ lv y +d<→lv Φ< = refl +d<→lv (s< lt) = refl + +o∅ : {n : Level} → Ordinal {n} +o∅ = record { lv = Zero ; ord = Φ Zero } + +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 + +≡→¬d< : {n : Level} → {lv : Nat} → {x : OrdinalD {n} lv } → x d< x → ⊥ +≡→¬d< {n} {lx} {OSuc lx y} (s< t) = ≡→¬d< t + +trio<≡ : {n : Level} → {lx : Nat} {x : OrdinalD {n} lx } { y : OrdinalD lx } → x ≡ y → x d< y → ⊥ +trio<≡ refl = ≡→¬d< + +trio>≡ : {n : Level} → {lx : Nat} {x : OrdinalD {n} lx } { y : OrdinalD lx } → x ≡ y → y d< x → ⊥ +trio>≡ refl = ≡→¬d< + +triOrdd : {n : Level} → {lx : Nat} → Trichotomous _≡_ ( _d<_ {n} {lx} {lx} ) +triOrdd {_} {lv} (Φ lv) (Φ lv) = tri≈ ≡→¬d< refl ≡→¬d< +triOrdd {_} {lv} (Φ lv) (OSuc lv y) = tri< Φ< (λ ()) ( λ lt → trio<> lt Φ< ) +triOrdd {_} {lv} (OSuc lv x) (Φ lv) = tri> (λ lt → trio<> lt Φ<) (λ ()) Φ< +triOrdd {_} {lv} (OSuc lv x) (OSuc lv y) with triOrdd x y +triOrdd {_} {lv} (OSuc lv x) (OSuc lv y) | tri< a ¬b ¬c = tri< (s< a) (λ tx=ty → trio<≡ tx=ty (s< a) ) ( λ lt → trio<> lt (s< a) ) +triOrdd {_} {lv} (OSuc lv x) (OSuc lv x) | tri≈ ¬a refl ¬c = tri≈ ≡→¬d< refl ≡→¬d< +triOrdd {_} {lv} (OSuc lv x) (OSuc lv y) | tri> ¬a ¬b c = tri> ( λ lt → trio<> lt (s< c) ) (λ tx=ty → trio>≡ tx=ty (s< c) ) (s< c) + +osuc : {n : Level} ( x : Ordinal {n} ) → Ordinal {n} +osuc record { lv = lx ; ord = ox } = record { lv = lx ; ord = OSuc lx ox } + +<-osuc : {n : Level} { x : Ordinal {n} } → x o< osuc x +<-osuc {n} {record { lv = lx ; ord = Φ .lx }} = case2 Φ< +<-osuc {n} {record { lv = lx ; ord = OSuc .lx ox }} = case2 ( s< s<refl ) + +o<¬≡ : {n : Level } { ox oy : Ordinal {suc n}} → ox ≡ oy → ox o< oy → ⊥ +o<¬≡ {_} {ox} {ox} refl (case1 lt) = =→¬< lt +o<¬≡ {_} {ox} {ox} refl (case2 (s< lt)) = trio<≡ refl lt + +¬x<0 : {n : Level} → { x : Ordinal {suc n} } → ¬ ( x o< o∅ {suc n} ) +¬x<0 {n} {x} (case1 ()) +¬x<0 {n} {x} (case2 ()) + +o<> : {n : Level} → {x y : Ordinal {n} } → y o< x → x o< y → ⊥ +o<> {n} {x} {y} (case1 x₁) (case1 x₂) = nat-<> x₁ x₂ +o<> {n} {x} {y} (case1 x₁) (case2 x₂) = nat-≡< (sym (d<→lv x₂)) x₁ +o<> {n} {x} {y} (case2 x₁) (case1 x₂) = nat-≡< (sym (d<→lv x₁)) x₂ +o<> {n} {record { lv = lv₁ ; ord = .(OSuc lv₁ _) }} {record { lv = .lv₁ ; ord = .(Φ lv₁) }} (case2 Φ<) (case2 ()) +o<> {n} {record { lv = lv₁ ; ord = .(OSuc lv₁ _) }} {record { lv = .lv₁ ; ord = .(OSuc lv₁ _) }} (case2 (s< y<x)) (case2 (s< x<y)) = + o<> (case2 y<x) (case2 x<y) + +orddtrans : {n : Level} {lx : Nat} {x y z : OrdinalD {n} lx } → x d< y → y d< z → x d< z +orddtrans {_} {lx} {.(Φ lx)} {.(OSuc lx _)} {.(OSuc lx _)} Φ< (s< y<z) = Φ< +orddtrans {_} {lx} {.(OSuc lx _)} {.(OSuc lx _)} {.(OSuc lx _)} (s< x<y) (s< y<z) = s< ( orddtrans x<y y<z ) + +osuc-≡< : {n : Level} { a x : Ordinal {n} } → x o< osuc a → (x ≡ a ) ∨ (x o< a) +osuc-≡< {n} {a} {x} (case1 lt) = case2 (case1 lt) +osuc-≡< {n} {record { lv = lv₁ ; ord = Φ .lv₁ }} {record { lv = .lv₁ ; ord = .(Φ lv₁) }} (case2 Φ<) = case1 refl +osuc-≡< {n} {record { lv = lv₁ ; ord = OSuc .lv₁ ord₁ }} {record { lv = .lv₁ ; ord = .(Φ lv₁) }} (case2 Φ<) = case2 (case2 Φ<) +osuc-≡< {n} {record { lv = lv₁ ; ord = Φ .lv₁ }} {record { lv = .lv₁ ; ord = .(OSuc lv₁ _) }} (case2 (s< ())) +osuc-≡< {n} {record { lv = la ; ord = OSuc la oa }} {record { lv = la ; ord = (OSuc la ox) }} (case2 (s< lt)) with + osuc-≡< {n} {record { lv = la ; ord = oa }} {record { lv = la ; ord = ox }} (case2 lt ) +... | case1 refl = case1 refl +... | case2 (case2 x) = case2 (case2( s< x) ) +... | case2 (case1 x) = ⊥-elim (¬a≤a x) + +osuc-< : {n : Level} { x y : Ordinal {n} } → y o< osuc x → x o< y → ⊥ +osuc-< {n} {x} {y} y<ox x<y with osuc-≡< y<ox +osuc-< {n} {x} {x} y<ox (case1 x₁) | case1 refl = ⊥-elim (¬a≤a x₁) +osuc-< {n} {x} {x} (case1 x₂) (case2 x₁) | case1 refl = ⊥-elim (¬a≤a x₂) +osuc-< {n} {x} {x} (case2 x₂) (case2 x₁) | case1 refl = ≡→¬d< x₁ +osuc-< {n} {x} {y} y<ox (case1 x₂) | case2 (case1 x₁) = nat-<> x₂ x₁ +osuc-< {n} {x} {y} y<ox (case1 x₂) | case2 (case2 x₁) = nat-≡< (sym (d<→lv x₁)) x₂ +osuc-< {n} {x} {y} y<ox (case2 x<y) | case2 y<x = o<> (case2 x<y) y<x + + +ordtrans : {n : Level} {x y z : Ordinal {n} } → x o< y → y o< z → x o< z +ordtrans {n} {x} {y} {z} (case1 x₁) (case1 x₂) = case1 ( <-trans x₁ x₂ ) +ordtrans {n} {x} {y} {z} (case1 x₁) (case2 x₂) with d<→lv x₂ +... | refl = case1 x₁ +ordtrans {n} {x} {y} {z} (case2 x₁) (case1 x₂) with d<→lv x₁ +... | refl = case1 x₂ +ordtrans {n} {x} {y} {z} (case2 x₁) (case2 x₂) with d<→lv x₁ | d<→lv x₂ +... | refl | refl = case2 ( orddtrans x₁ x₂ ) + +trio< : {n : Level } → Trichotomous {suc n} _≡_ _o<_ +trio< a b with <-cmp (lv a) (lv b) +trio< a b | tri< a₁ ¬b ¬c = tri< (case1 a₁) (λ refl → ¬b (cong ( λ x → lv x ) refl ) ) lemma1 where + lemma1 : ¬ (Suc (lv b) ≤ lv a) ∨ (ord b d< ord a) + lemma1 (case1 x) = ¬c x + lemma1 (case2 x) = ⊥-elim (nat-≡< (sym ( d<→lv x )) a₁ ) +trio< a b | tri> ¬a ¬b c = tri> lemma1 (λ refl → ¬b (cong ( λ x → lv x ) refl ) ) (case1 c) where + lemma1 : ¬ (Suc (lv a) ≤ lv b) ∨ (ord a d< ord b) + lemma1 (case1 x) = ¬a x + lemma1 (case2 x) = ⊥-elim (nat-≡< (sym ( d<→lv x )) c ) +trio< a b | tri≈ ¬a refl ¬c with triOrdd ( ord a ) ( ord b ) +trio< record { lv = .(lv b) ; ord = x } b | tri≈ ¬a refl ¬c | tri< a ¬b ¬c₁ = tri< (case2 a) (λ refl → ¬b (lemma1 refl )) lemma2 where + lemma1 : (record { lv = _ ; ord = x }) ≡ b → x ≡ ord b + lemma1 refl = refl + lemma2 : ¬ (Suc (lv b) ≤ lv b) ∨ (ord b d< x) + lemma2 (case1 x) = ¬a x + lemma2 (case2 x) = trio<> x a +trio< record { lv = .(lv b) ; ord = x } b | tri≈ ¬a refl ¬c | tri> ¬a₁ ¬b c = tri> lemma2 (λ refl → ¬b (lemma1 refl )) (case2 c) where + lemma1 : (record { lv = _ ; ord = x }) ≡ b → x ≡ ord b + lemma1 refl = refl + lemma2 : ¬ (Suc (lv b) ≤ lv b) ∨ (x d< ord b) + lemma2 (case1 x) = ¬a x + lemma2 (case2 x) = trio<> x c +trio< record { lv = .(lv b) ; ord = x } b | tri≈ ¬a refl ¬c | tri≈ ¬a₁ refl ¬c₁ = tri≈ lemma1 refl lemma1 where + lemma1 : ¬ (Suc (lv b) ≤ lv b) ∨ (ord b d< ord b) + lemma1 (case1 x) = ¬a x + lemma1 (case2 x) = ≡→¬d< x + + +open _∧_ + +TransFinite : {n m : Level} → { ψ : Ordinal {suc n} → Set m } + → ( ∀ (lx : Nat ) → ( (x : Ordinal {suc n} ) → x o< ordinal lx (Φ lx) → ψ x ) → ψ ( record { lv = lx ; ord = Φ lx } ) ) + → ( ∀ (lx : Nat ) → (x : OrdinalD lx ) → ( (y : Ordinal {suc n} ) → y o< ordinal lx (OSuc lx x) → ψ y ) → ψ ( record { lv = lx ; ord = OSuc lx x } ) ) + → ∀ (x : Ordinal) → ψ x +TransFinite {n} {m} {ψ} caseΦ caseOSuc x = proj1 (TransFinite1 (lv x) (ord x) ) where + TransFinite1 : (lx : Nat) (ox : OrdinalD lx ) → ψ (ordinal lx ox) ∧ ( ( (x : Ordinal {suc n} ) → x o< ordinal lx ox → ψ x ) ) + TransFinite1 Zero (Φ 0) = ⟪ caseΦ Zero lemma , lemma1 ⟫ where + lemma : (x : Ordinal) → x o< ordinal Zero (Φ Zero) → ψ x + lemma x (case1 ()) + lemma x (case2 ()) + lemma1 : (x : Ordinal) → x o< ordinal Zero (Φ Zero) → ψ x + lemma1 x (case1 ()) + lemma1 x (case2 ()) + TransFinite1 (Suc lx) (Φ (Suc lx)) = ⟪ caseΦ (Suc lx) (λ x → lemma (lv x) (ord x)) , (λ x → lemma (lv x) (ord x)) ⟫ where + lemma0 : (ly : Nat) (oy : OrdinalD ly ) → ordinal ly oy o< ordinal lx (Φ lx) → ψ (ordinal ly oy) + lemma0 ly oy lt = proj2 ( TransFinite1 lx (Φ lx) ) (ordinal ly oy) lt + lemma : (ly : Nat) (oy : OrdinalD ly ) → ordinal ly oy o< ordinal (Suc lx) (Φ (Suc lx)) → ψ (ordinal ly oy) + lemma lx1 ox1 (case1 lt) with <-∨ lt + lemma lx (Φ lx) (case1 lt) | case1 refl = proj1 ( TransFinite1 lx (Φ lx) ) + lemma lx (Φ lx) (case1 lt) | case2 lt1 = lemma0 lx (Φ lx) (case1 lt1) + lemma lx (OSuc lx ox1) (case1 lt) | case1 refl = caseOSuc lx ox1 lemma2 where + lemma2 : (y : Ordinal) → (Suc (lv y) ≤ lx) ∨ (ord y d< OSuc lx ox1) → ψ y + lemma2 y lt1 with osuc-≡< lt1 + lemma2 y lt1 | case1 refl = lemma lx ox1 (case1 a<sa) + lemma2 y lt1 | case2 t = proj2 (TransFinite1 lx ox1) y t + lemma lx1 (OSuc lx1 ox1) (case1 lt) | case2 lt1 = caseOSuc lx1 ox1 lemma2 where + lemma2 : (y : Ordinal) → (Suc (lv y) ≤ lx1) ∨ (ord y d< OSuc lx1 ox1) → ψ y + lemma2 y lt2 with osuc-≡< lt2 + lemma2 y lt2 | case1 refl = lemma lx1 ox1 (ordtrans lt2 (case1 lt)) + lemma2 y lt2 | case2 (case1 lt3) = proj2 (TransFinite1 lx (Φ lx)) y (case1 (<-trans lt3 lt1 )) + lemma2 y lt2 | case2 (case2 lt3) with d<→lv lt3 + ... | refl = proj2 (TransFinite1 lx (Φ lx)) y (case1 lt1) + TransFinite1 lx (OSuc lx ox) = ⟪ caseOSuc lx ox lemma , lemma ⟫ where + lemma : (y : Ordinal) → y o< ordinal lx (OSuc lx ox) → ψ y + lemma y lt with osuc-≡< lt + lemma y lt | case1 refl = proj1 ( TransFinite1 lx ox ) + lemma y lt | case2 lt1 = proj2 ( TransFinite1 lx ox ) y lt1 + +-- record CountableOrdinal {n : Level} : Set (suc (suc n)) where +-- field +-- ctl→ : Nat → Ordinal {suc n} +-- ctl← : Ordinal → Nat +-- ctl-iso→ : { x : Ordinal } → ctl→ (ctl← x ) ≡ x +-- ctl-iso← : { x : Nat } → ctl← (ctl→ x ) ≡ x +-- +-- is-C-Ordinal : {n : Level} → CountableOrdinal {n} +-- is-C-Ordinal {n} = record { +-- ctl→ = {!!} +-- ; ctl← = λ x → TransFinite {n} (λ lx lt → Zero ) ctl01 x +-- ; ctl-iso→ = {!!} +-- ; ctl-iso← = {!!} +-- } where +-- ctl01 : (lx : Nat) (x : OrdinalD lx) → ((y : Ordinal) → y o< ordinal lx (OSuc lx x) → Nat) → Nat +-- ctl01 Zero (Φ Zero) prev = Zero +-- ctl01 Zero (OSuc Zero x) prev = Suc ( prev (ordinal Zero x) (ordtrans <-osuc <-osuc )) +-- ctl01 (Suc lx) (Φ (Suc lx)) prev = Suc ( prev (ordinal lx {!!}) {!!}) +-- ctl01 (Suc lx) (OSuc (Suc lx) x) prev = Suc ( prev (ordinal (Suc lx) x) (ordtrans <-osuc <-osuc )) + +open import Ordinals + +C-Ordinal : {n : Level} → Ordinals {suc n} +C-Ordinal {n} = record { + Ordinal = Ordinal {suc n} + ; o∅ = o∅ + ; osuc = osuc + ; _o<_ = _o<_ + ; next = next + ; isOrdinal = record { + ordtrans = ordtrans + ; trio< = trio< + ; ¬x<0 = ¬x<0 + ; <-osuc = <-osuc + ; osuc-≡< = osuc-≡< + ; TransFinite = TransFinite2 + ; Oprev-p = Oprev-p + } ; + isNext = record { + x<nx = x<nx + ; osuc<nx = λ {x} {y} → osuc<nx {x} {y} + ; ¬nx<nx = ¬nx<nx + } + } where + next : Ordinal {suc n} → Ordinal {suc n} + next (ordinal lv ord) = ordinal (Suc lv) (Φ (Suc lv)) + x<nx : { y : Ordinal } → (y o< next y ) + x<nx = case1 a<sa + osuc<nx : { x y : Ordinal } → x o< next y → osuc x o< next y + osuc<nx (case1 lt) = case1 lt + ¬nx<nx : {x y : Ordinal} → y o< x → x o< next y → ¬ ((z : Ordinal) → ¬ (x ≡ osuc z)) + ¬nx<nx {x} {y} = lemma2 x where + lemma2 : (x : Ordinal) → y o< x → x o< next y → ¬ ((z : Ordinal) → ¬ x ≡ osuc z) + lemma2 (ordinal Zero (Φ 0)) (case2 ()) (case1 (s≤s z≤n)) not + lemma2 (ordinal Zero (OSuc 0 dx)) (case2 Φ<) (case1 (s≤s z≤n)) not = not _ refl + lemma2 (ordinal Zero (OSuc 0 dx)) (case2 (s< x)) (case1 (s≤s z≤n)) not = not _ refl + lemma2 (ordinal (Suc lx) (OSuc (Suc lx) ox)) y<x (case1 (s≤s (s≤s lt))) not = not _ refl + lemma2 (ordinal (Suc lx) (Φ (Suc lx))) (case1 x) (case1 (s≤s (s≤s lt))) not = lemma3 x lt where + lemma3 : {n l : Nat} → (Suc (Suc n) ≤ Suc l) → l ≤ n → ⊥ + lemma3 (s≤s sn≤l) (s≤s l≤n) = lemma3 sn≤l l≤n + open Oprev + Oprev-p : (x : Ordinal) → Dec ( Oprev (Ordinal {suc n}) osuc x ) + Oprev-p (ordinal lv (Φ lv)) = no (λ not → lemma (oprev not) (oprev=x not) ) where + lemma : (x : Ordinal) → osuc x ≡ (ordinal lv (Φ lv)) → ⊥ + lemma x () + Oprev-p (ordinal lv (OSuc lv ox)) = yes record { oprev = ordinal lv ox ; oprev=x = refl } + ord1 : Set (suc n) + ord1 = Ordinal {suc n} + TransFinite2 : { ψ : ord1 → Set (suc (suc n)) } + → ( (x : ord1) → ( (y : ord1 ) → y o< x → ψ y ) → ψ x ) + → ∀ (x : ord1) → ψ x + TransFinite2 {ψ} lt x = TransFinite {n} {suc (suc n)} {ψ} caseΦ caseOSuc x where + caseΦ : (lx : Nat) → ((x₁ : Ordinal) → x₁ o< ordinal lx (Φ lx) → ψ x₁) → + ψ (record { lv = lx ; ord = Φ lx }) + caseΦ lx prev = lt (ordinal lx (Φ lx) ) prev + caseOSuc : (lx : Nat) (x₁ : OrdinalD lx) → ((y : Ordinal) → y o< ordinal lx (OSuc lx x₁) → ψ y) → + ψ (record { lv = lx ; ord = OSuc lx x₁ }) + caseOSuc lx ox prev = lt (ordinal lx (OSuc lx ox)) prev + +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/partfunc.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,227 @@ +{-# OPTIONS --allow-unsolved-metas #-} +open import Level +open import Relation.Nullary +open import Relation.Binary.PropositionalEquality +open import Ordinals +module partfunc {n : Level } (O : Ordinals {n}) where + +open import logic +open import Relation.Binary +open import Data.Empty +open import Data.List hiding (filter) +open import Data.Maybe +open import Relation.Binary +open import Relation.Binary.Core +open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) +open import filter O + +open _∧_ +open _∨_ +open Bool + +---- +-- +-- Partial Function without ZF +-- + +record PFunc (Dom : Set n) (Cod : Set n) : Set (suc n) where + field + dom : Dom → Set n + pmap : (x : Dom ) → dom x → Cod + meq : {x : Dom } → { p q : dom x } → pmap x p ≡ pmap x q + +---- +-- +-- PFunc (Lift n Nat) Cod is equivalent to List (Maybe Cod) +-- + +data Findp : {Cod : Set n} → List (Maybe Cod) → (x : Nat) → Set where + v0 : {Cod : Set n} → {f : List (Maybe Cod)} → ( v : Cod ) → Findp ( just v ∷ f ) Zero + vn : {Cod : Set n} → {f : List (Maybe Cod)} {d : Maybe Cod} → {x : Nat} → Findp f x → Findp (d ∷ f) (Suc x) + +open PFunc + +find : {Cod : Set n} → (f : List (Maybe Cod) ) → (x : Nat) → Findp f x → Cod +find (just v ∷ _) 0 (v0 v) = v +find (_ ∷ n) (Suc i) (vn p) = find n i p + +findpeq : {Cod : Set n} → (f : List (Maybe Cod)) → {x : Nat} {p q : Findp f x } → find f x p ≡ find f x q +findpeq n {0} {v0 _} {v0 _} = refl +findpeq [] {Suc x} {()} +findpeq (just x₁ ∷ n) {Suc x} {vn p} {vn q} = findpeq n {x} {p} {q} +findpeq (nothing ∷ n) {Suc x} {vn p} {vn q} = findpeq n {x} {p} {q} + +List→PFunc : {Cod : Set n} → List (Maybe Cod) → PFunc (Lift n Nat) Cod +List→PFunc fp = record { dom = λ x → Lift n (Findp fp (lower x)) + ; pmap = λ x y → find fp (lower x) (lower y) + ; meq = λ {x} {p} {q} → findpeq fp {lower x} {lower p} {lower q} + } +---- +-- +-- to List (Maybe Two) is a Latice +-- + +_3⊆b_ : (f g : List (Maybe Two)) → Bool +[] 3⊆b [] = true +[] 3⊆b (nothing ∷ g) = [] 3⊆b g +[] 3⊆b (_ ∷ g) = true +(nothing ∷ f) 3⊆b [] = f 3⊆b [] +(nothing ∷ f) 3⊆b (_ ∷ g) = f 3⊆b g +(just i0 ∷ f) 3⊆b (just i0 ∷ g) = f 3⊆b g +(just i1 ∷ f) 3⊆b (just i1 ∷ g) = f 3⊆b g +_ 3⊆b _ = false + +_3⊆_ : (f g : List (Maybe Two)) → Set +f 3⊆ g = (f 3⊆b g) ≡ true + +_3∩_ : (f g : List (Maybe Two)) → List (Maybe Two) +[] 3∩ (nothing ∷ g) = nothing ∷ ([] 3∩ g) +[] 3∩ g = [] +(nothing ∷ f) 3∩ [] = nothing ∷ f 3∩ [] +f 3∩ [] = [] +(just i0 ∷ f) 3∩ (just i0 ∷ g) = just i0 ∷ ( f 3∩ g ) +(just i1 ∷ f) 3∩ (just i1 ∷ g) = just i1 ∷ ( f 3∩ g ) +(_ ∷ f) 3∩ (_ ∷ g) = nothing ∷ ( f 3∩ g ) + +3∩⊆f : { f g : List (Maybe Two) } → (f 3∩ g ) 3⊆ f +3∩⊆f {[]} {[]} = refl +3∩⊆f {[]} {just _ ∷ g} = refl +3∩⊆f {[]} {nothing ∷ g} = 3∩⊆f {[]} {g} +3∩⊆f {just _ ∷ f} {[]} = refl +3∩⊆f {nothing ∷ f} {[]} = 3∩⊆f {f} {[]} +3∩⊆f {just i0 ∷ f} {just i0 ∷ g} = 3∩⊆f {f} {g} +3∩⊆f {just i1 ∷ f} {just i1 ∷ g} = 3∩⊆f {f} {g} +3∩⊆f {just i0 ∷ f} {just i1 ∷ g} = 3∩⊆f {f} {g} +3∩⊆f {just i1 ∷ f} {just i0 ∷ g} = 3∩⊆f {f} {g} +3∩⊆f {nothing ∷ f} {just _ ∷ g} = 3∩⊆f {f} {g} +3∩⊆f {just i0 ∷ f} {nothing ∷ g} = 3∩⊆f {f} {g} +3∩⊆f {just i1 ∷ f} {nothing ∷ g} = 3∩⊆f {f} {g} +3∩⊆f {nothing ∷ f} {nothing ∷ g} = 3∩⊆f {f} {g} + +3∩sym : { f g : List (Maybe Two) } → (f 3∩ g ) ≡ (g 3∩ f ) +3∩sym {[]} {[]} = refl +3∩sym {[]} {just _ ∷ g} = refl +3∩sym {[]} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {[]} {g}) +3∩sym {just _ ∷ f} {[]} = refl +3∩sym {nothing ∷ f} {[]} = cong (λ k → nothing ∷ k) (3∩sym {f} {[]}) +3∩sym {just i0 ∷ f} {just i0 ∷ g} = cong (λ k → just i0 ∷ k) (3∩sym {f} {g}) +3∩sym {just i0 ∷ f} {just i1 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) +3∩sym {just i1 ∷ f} {just i0 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) +3∩sym {just i1 ∷ f} {just i1 ∷ g} = cong (λ k → just i1 ∷ k) (3∩sym {f} {g}) +3∩sym {just i0 ∷ f} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) +3∩sym {just i1 ∷ f} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) +3∩sym {nothing ∷ f} {just i0 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) +3∩sym {nothing ∷ f} {just i1 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) +3∩sym {nothing ∷ f} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g}) + +3⊆-[] : { h : List (Maybe Two) } → [] 3⊆ h +3⊆-[] {[]} = refl +3⊆-[] {just _ ∷ h} = refl +3⊆-[] {nothing ∷ h} = 3⊆-[] {h} + +3⊆trans : { f g h : List (Maybe Two) } → f 3⊆ g → g 3⊆ h → f 3⊆ h +3⊆trans {[]} {[]} {[]} f<g g<h = refl +3⊆trans {[]} {[]} {just _ ∷ h} f<g g<h = refl +3⊆trans {[]} {[]} {nothing ∷ h} f<g g<h = 3⊆trans {[]} {[]} {h} refl g<h +3⊆trans {[]} {nothing ∷ g} {[]} f<g g<h = refl +3⊆trans {[]} {just _ ∷ g} {just _ ∷ h} f<g g<h = refl +3⊆trans {[]} {nothing ∷ g} {just _ ∷ h} f<g g<h = refl +3⊆trans {[]} {nothing ∷ g} {nothing ∷ h} f<g g<h = 3⊆trans {[]} {g} {h} f<g g<h +3⊆trans {nothing ∷ f} {[]} {[]} f<g g<h = f<g +3⊆trans {nothing ∷ f} {[]} {just _ ∷ h} f<g g<h = 3⊆trans {f} {[]} {h} f<g (3⊆-[] {h}) +3⊆trans {nothing ∷ f} {[]} {nothing ∷ h} f<g g<h = 3⊆trans {f} {[]} {h} f<g g<h +3⊆trans {nothing ∷ f} {nothing ∷ g} {[]} f<g g<h = 3⊆trans {f} {g} {[]} f<g g<h +3⊆trans {nothing ∷ f} {nothing ∷ g} {just _ ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h +3⊆trans {nothing ∷ f} {nothing ∷ g} {nothing ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h +3⊆trans {[]} {just i0 ∷ g} {[]} f<g () +3⊆trans {[]} {just i1 ∷ g} {[]} f<g () +3⊆trans {[]} {just x ∷ g} {nothing ∷ h} f<g g<h = 3⊆-[] {h} +3⊆trans {just i0 ∷ f} {[]} {h} () g<h +3⊆trans {just i1 ∷ f} {[]} {h} () g<h +3⊆trans {just x ∷ f} {just i0 ∷ g} {[]} f<g () +3⊆trans {just x ∷ f} {just i1 ∷ g} {[]} f<g () +3⊆trans {just i0 ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h +3⊆trans {just i1 ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h +3⊆trans {just x ∷ f} {just i0 ∷ g} {nothing ∷ h} f<g () +3⊆trans {just x ∷ f} {just i1 ∷ g} {nothing ∷ h} f<g () +3⊆trans {just i0 ∷ f} {nothing ∷ g} {_} () g<h +3⊆trans {just i1 ∷ f} {nothing ∷ g} {_} () g<h +3⊆trans {nothing ∷ f} {just i0 ∷ g} {[]} f<g () +3⊆trans {nothing ∷ f} {just i1 ∷ g} {[]} f<g () +3⊆trans {nothing ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h +3⊆trans {nothing ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h +3⊆trans {nothing ∷ f} {just i0 ∷ g} {nothing ∷ h} f<g () +3⊆trans {nothing ∷ f} {just i1 ∷ g} {nothing ∷ h} f<g () + +3⊆∩f : { f g h : List (Maybe Two) } → f 3⊆ g → f 3⊆ h → f 3⊆ (g 3∩ h ) +3⊆∩f {[]} {[]} {[]} f<g f<h = refl +3⊆∩f {[]} {[]} {x ∷ h} f<g f<h = 3⊆-[] {[] 3∩ (x ∷ h)} +3⊆∩f {[]} {x ∷ g} {h} f<g f<h = 3⊆-[] {(x ∷ g) 3∩ h} +3⊆∩f {nothing ∷ f} {[]} {[]} f<g f<h = 3⊆∩f {f} {[]} {[]} f<g f<h +3⊆∩f {nothing ∷ f} {[]} {just _ ∷ h} f<g f<h = f<g +3⊆∩f {nothing ∷ f} {[]} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {[]} {h} f<g f<h +3⊆∩f {just i0 ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h +3⊆∩f {just i1 ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h +3⊆∩f {nothing ∷ f} {just _ ∷ g} {[]} f<g f<h = f<h +3⊆∩f {nothing ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h +3⊆∩f {nothing ∷ f} {just i0 ∷ g} {just i1 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h +3⊆∩f {nothing ∷ f} {just i1 ∷ g} {just i0 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h +3⊆∩f {nothing ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h +3⊆∩f {nothing ∷ f} {just i0 ∷ g} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h +3⊆∩f {nothing ∷ f} {just i1 ∷ g} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h +3⊆∩f {nothing ∷ f} {nothing ∷ g} {[]} f<g f<h = 3⊆∩f {f} {g} {[]} f<g f<h +3⊆∩f {nothing ∷ f} {nothing ∷ g} {just _ ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h +3⊆∩f {nothing ∷ f} {nothing ∷ g} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h + +3↑22 : (f : Nat → Two) (i j : Nat) → List (Maybe Two) +3↑22 f Zero j = [] +3↑22 f (Suc i) j = just (f j) ∷ 3↑22 f i (Suc j) + +_3↑_ : (Nat → Two) → Nat → List (Maybe Two) +_3↑_ f i = 3↑22 f i 0 + +3↑< : {f : Nat → Two} → { x y : Nat } → x ≤ y → (_3↑_ f x) 3⊆ (_3↑_ f y) +3↑< {f} {x} {y} x<y = lemma x y 0 x<y where + lemma : (x y i : Nat) → x ≤ y → (3↑22 f x i ) 3⊆ (3↑22 f y i ) + lemma 0 y i z≤n with f i + lemma Zero Zero i z≤n | i0 = refl + lemma Zero (Suc y) i z≤n | i0 = 3⊆-[] {3↑22 f (Suc y) i} + lemma Zero Zero i z≤n | i1 = refl + lemma Zero (Suc y) i z≤n | i1 = 3⊆-[] {3↑22 f (Suc y) i} + lemma (Suc x) (Suc y) i (s≤s x<y) with f i + lemma (Suc x) (Suc y) i (s≤s x<y) | i0 = lemma x y (Suc i) x<y + lemma (Suc x) (Suc y) i (s≤s x<y) | i1 = lemma x y (Suc i) x<y + +Finite3b : (p : List (Maybe Two) ) → Bool +Finite3b [] = true +Finite3b (just _ ∷ f) = Finite3b f +Finite3b (nothing ∷ f) = false + +finite3cov : (p : List (Maybe Two) ) → List (Maybe Two) +finite3cov [] = [] +finite3cov (just y ∷ x) = just y ∷ finite3cov x +finite3cov (nothing ∷ x) = just i0 ∷ finite3cov x + +Dense-3 : F-Dense (List (Maybe Two) ) (λ x → One) _3⊆_ _3∩_ +Dense-3 = record { + dense = λ x → Finite3b x ≡ true + ; d⊆P = OneObj + ; dense-f = λ x → finite3cov x + ; dense-d = λ {p} d → lemma1 p + ; dense-p = λ {p} d → lemma2 p + } where + lemma1 : (p : List (Maybe Two) ) → Finite3b (finite3cov p) ≡ true + lemma1 [] = refl + lemma1 (just i0 ∷ p) = lemma1 p + lemma1 (just i1 ∷ p) = lemma1 p + lemma1 (nothing ∷ p) = lemma1 p + lemma2 : (p : List (Maybe Two)) → p 3⊆ finite3cov p + lemma2 [] = refl + lemma2 (just i0 ∷ p) = lemma2 p + lemma2 (just i1 ∷ p) = lemma2 p + lemma2 (nothing ∷ p) = lemma2 p + +-- min = Data.Nat._⊓_ +-- m≤m⊔n = Data.Nat._⊔_ +-- open import Data.Nat.Properties +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/zf.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,80 @@ +module zf where + +open import Level + +open import logic + +open import Relation.Nullary +open import Relation.Binary hiding (_⇔_) +open import Data.Empty + +record IsZF {n m : Level } + (ZFSet : Set n) + (_∋_ : ( A x : ZFSet ) → Set m) + (_≈_ : Rel ZFSet m) + (∅ : ZFSet) + (_,_ : ( A B : ZFSet ) → ZFSet) + (Union : ( A : ZFSet ) → ZFSet) + (Power : ( A : ZFSet ) → ZFSet) + (Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet ) + (Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet ) + (infinite : ZFSet) + : Set (suc (n ⊔ suc m)) where + field + isEquivalence : IsEquivalence {n} {m} {ZFSet} _≈_ + -- ∀ x ∀ y ∃ z ∀ t ( z ∋ t → t ≈ x ∨ t ≈ y) + pair→ : ( x y t : ZFSet ) → (x , y) ∋ t → ( t ≈ x ) ∨ ( t ≈ y ) + pair← : ( x y t : ZFSet ) → ( t ≈ x ) ∨ ( t ≈ y ) → (x , y) ∋ t + -- ∀ x ∃ y ∀ z (z ∈ y ⇔ ∃ u ∈ x ∧ (z ∈ u)) + union→ : ( X z u : ZFSet ) → ( X ∋ u ) ∧ (u ∋ z ) → Union X ∋ z + union← : ( X z : ZFSet ) → (X∋z : Union X ∋ z ) → ¬ ( (u : ZFSet ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) + _∈_ : ( A B : ZFSet ) → Set m + A ∈ B = B ∋ A + _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set m + _⊆_ A B {x} = A ∋ x → B ∋ x + _∩_ : ( A B : ZFSet ) → ZFSet + A ∩ B = Select A ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) + _∪_ : ( A B : ZFSet ) → ZFSet + A ∪ B = Union (A , B) + {_} : ZFSet → ZFSet + { x } = ( x , x ) + infixr 200 _∈_ + infixr 230 _∩_ _∪_ + infixr 220 _⊆_ + field + empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x ) + -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ t ⊆ X ) ) + power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → t ∋ x → ¬ ¬ ( A ∋ x ) -- _⊆_ t A {x} + power← : ∀( A t : ZFSet ) → ( ∀ {x} → _⊆_ t A {x}) → Power A ∋ t + -- extensionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w ) + extensionality : { A B w : ZFSet } → ( (z : ZFSet) → ( A ∋ z ) ⇔ (B ∋ z) ) → ( A ∈ w ⇔ B ∈ w ) + -- regularity without minimum + ε-induction : { ψ : ZFSet → Set (suc m)} + → ( {x : ZFSet } → ({ y : ZFSet } → x ∋ y → ψ y ) → ψ x ) + → (x : ZFSet ) → ψ x + -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) ) + infinity∅ : ∅ ∈ infinite + infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ { x }) ∈ infinite + selection : { ψ : ZFSet → Set m } → ∀ { X y : ZFSet } → ( ( y ∈ X ) ∧ ψ y ) ⇔ (y ∈ Select X ψ ) + -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) ) + replacement← : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → x ∈ X → ψ x ∈ Replace X ψ + replacement→ : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → ( lt : x ∈ Replace X ψ ) → ¬ ( ∀ (y : ZFSet) → ¬ ( x ≈ ψ y ) ) + +record ZF {n m : Level } : Set (suc (n ⊔ suc m)) where + infixr 210 _,_ + infixl 200 _∋_ + infixr 220 _≈_ + field + ZFSet : Set n + _∋_ : ( A x : ZFSet ) → Set m + _≈_ : ( A B : ZFSet ) → Set m + -- ZF Set constructor + ∅ : ZFSet + _,_ : ( A B : ZFSet ) → ZFSet + Union : ( A : ZFSet ) → ZFSet + Power : ( A : ZFSet ) → ZFSet + Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet + Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet + infinite : ZFSet + isZF : IsZF ZFSet _∋_ _≈_ ∅ _,_ Union Power Select Replace infinite +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/zfc.agda Mon Dec 21 10:23:37 2020 +0900 @@ -0,0 +1,34 @@ +module zfc where + +open import Level +open import Relation.Binary +open import Relation.Nullary +open import logic + +record IsZFC {n m : Level } + (ZFSet : Set n) + (_∋_ : ( A x : ZFSet ) → Set m) + (_≈_ : Rel ZFSet m) + (∅ : ZFSet) + (Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet ) + : Set (suc (n ⊔ suc m)) where + field + -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] + choice-func : (X : ZFSet ) → {x : ZFSet } → ¬ ( x ≈ ∅ ) → ( X ∋ x ) → ZFSet + choice : (X : ZFSet ) → {A : ZFSet } → ( X∋A : X ∋ A ) → (not : ¬ ( A ≈ ∅ )) → A ∋ choice-func X not X∋A + infixr 200 _∈_ + infixr 230 _∩_ + _∈_ : ( A B : ZFSet ) → Set m + A ∈ B = B ∋ A + _∩_ : ( A B : ZFSet ) → ZFSet + A ∩ B = Select A ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) + +record ZFC {n m : Level } : Set (suc (n ⊔ suc m)) where + field + ZFSet : Set n + _∋_ : ( A x : ZFSet ) → Set m + _≈_ : ( A B : ZFSet ) → Set m + ∅ : ZFSet + Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet + isZFC : IsZFC ZFSet _∋_ _≈_ ∅ Select +
--- a/zf.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,80 +0,0 @@ -module zf where - -open import Level - -open import logic - -open import Relation.Nullary -open import Relation.Binary hiding (_⇔_) -open import Data.Empty - -record IsZF {n m : Level } - (ZFSet : Set n) - (_∋_ : ( A x : ZFSet ) → Set m) - (_≈_ : Rel ZFSet m) - (∅ : ZFSet) - (_,_ : ( A B : ZFSet ) → ZFSet) - (Union : ( A : ZFSet ) → ZFSet) - (Power : ( A : ZFSet ) → ZFSet) - (Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet ) - (Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet ) - (infinite : ZFSet) - : Set (suc (n ⊔ suc m)) where - field - isEquivalence : IsEquivalence {n} {m} {ZFSet} _≈_ - -- ∀ x ∀ y ∃ z ∀ t ( z ∋ t → t ≈ x ∨ t ≈ y) - pair→ : ( x y t : ZFSet ) → (x , y) ∋ t → ( t ≈ x ) ∨ ( t ≈ y ) - pair← : ( x y t : ZFSet ) → ( t ≈ x ) ∨ ( t ≈ y ) → (x , y) ∋ t - -- ∀ x ∃ y ∀ z (z ∈ y ⇔ ∃ u ∈ x ∧ (z ∈ u)) - union→ : ( X z u : ZFSet ) → ( X ∋ u ) ∧ (u ∋ z ) → Union X ∋ z - union← : ( X z : ZFSet ) → (X∋z : Union X ∋ z ) → ¬ ( (u : ZFSet ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) - _∈_ : ( A B : ZFSet ) → Set m - A ∈ B = B ∋ A - _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set m - _⊆_ A B {x} = A ∋ x → B ∋ x - _∩_ : ( A B : ZFSet ) → ZFSet - A ∩ B = Select A ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) - _∪_ : ( A B : ZFSet ) → ZFSet - A ∪ B = Union (A , B) - {_} : ZFSet → ZFSet - { x } = ( x , x ) - infixr 200 _∈_ - infixr 230 _∩_ _∪_ - infixr 220 _⊆_ - field - empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x ) - -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ t ⊆ X ) ) - power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → t ∋ x → ¬ ¬ ( A ∋ x ) -- _⊆_ t A {x} - power← : ∀( A t : ZFSet ) → ( ∀ {x} → _⊆_ t A {x}) → Power A ∋ t - -- extensionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w ) - extensionality : { A B w : ZFSet } → ( (z : ZFSet) → ( A ∋ z ) ⇔ (B ∋ z) ) → ( A ∈ w ⇔ B ∈ w ) - -- regularity without minimum - ε-induction : { ψ : ZFSet → Set (suc m)} - → ( {x : ZFSet } → ({ y : ZFSet } → x ∋ y → ψ y ) → ψ x ) - → (x : ZFSet ) → ψ x - -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) ) - infinity∅ : ∅ ∈ infinite - infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ { x }) ∈ infinite - selection : { ψ : ZFSet → Set m } → ∀ { X y : ZFSet } → ( ( y ∈ X ) ∧ ψ y ) ⇔ (y ∈ Select X ψ ) - -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) ) - replacement← : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → x ∈ X → ψ x ∈ Replace X ψ - replacement→ : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → ( lt : x ∈ Replace X ψ ) → ¬ ( ∀ (y : ZFSet) → ¬ ( x ≈ ψ y ) ) - -record ZF {n m : Level } : Set (suc (n ⊔ suc m)) where - infixr 210 _,_ - infixl 200 _∋_ - infixr 220 _≈_ - field - ZFSet : Set n - _∋_ : ( A x : ZFSet ) → Set m - _≈_ : ( A B : ZFSet ) → Set m - -- ZF Set constructor - ∅ : ZFSet - _,_ : ( A B : ZFSet ) → ZFSet - Union : ( A : ZFSet ) → ZFSet - Power : ( A : ZFSet ) → ZFSet - Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet - Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet - infinite : ZFSet - isZF : IsZF ZFSet _∋_ _≈_ ∅ _,_ Union Power Select Replace infinite -
--- a/zfc.agda Sun Dec 20 12:37:07 2020 +0900 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,34 +0,0 @@ -module zfc where - -open import Level -open import Relation.Binary -open import Relation.Nullary -open import logic - -record IsZFC {n m : Level } - (ZFSet : Set n) - (_∋_ : ( A x : ZFSet ) → Set m) - (_≈_ : Rel ZFSet m) - (∅ : ZFSet) - (Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet ) - : Set (suc (n ⊔ suc m)) where - field - -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] - choice-func : (X : ZFSet ) → {x : ZFSet } → ¬ ( x ≈ ∅ ) → ( X ∋ x ) → ZFSet - choice : (X : ZFSet ) → {A : ZFSet } → ( X∋A : X ∋ A ) → (not : ¬ ( A ≈ ∅ )) → A ∋ choice-func X not X∋A - infixr 200 _∈_ - infixr 230 _∩_ - _∈_ : ( A B : ZFSet ) → Set m - A ∈ B = B ∋ A - _∩_ : ( A B : ZFSet ) → ZFSet - A ∩ B = Select A ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) - -record ZFC {n m : Level } : Set (suc (n ⊔ suc m)) where - field - ZFSet : Set n - _∋_ : ( A x : ZFSet ) → Set m - _≈_ : ( A B : ZFSet ) → Set m - ∅ : ZFSet - Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet - isZFC : IsZFC ZFSet _∋_ _≈_ ∅ Select -