annotate zf.agda @ 10:8022e14fce74

add constructible set
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Mon, 13 May 2019 18:25:38 +0900
parents 5ed16e2d8eb7
children 2df90eb0896c
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
1 module zf where
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
2
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
3 open import Level
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
4
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
5
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
6 record _∧_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
7 field
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
8 proj1 : A
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
9 proj2 : B
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
10
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
11 open _∧_
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
12
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
13
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
14 data _∨_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
15 case1 : A → A ∨ B
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
16 case2 : B → A ∨ B
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
17
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
18 -- open import Relation.Binary.PropositionalEquality
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
19
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
20 _⇔_ : {n : Level } → ( A B : Set n ) → Set n
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
21 _⇔_ A B = ( A → B ) ∧ ( B → A )
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
22
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
23 open import Data.Empty
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
24 open import Relation.Nullary
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
25
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
26 open import Relation.Binary
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
27 open import Relation.Binary.Core
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
28
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
29 infixr 130 _∧_
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
30 infixr 140 _∨_
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
31 infixr 150 _⇔_
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
32
9
5ed16e2d8eb7 try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 8
diff changeset
33 record Func {n m : Level } (ZFSet : Set n) (_≈_ : Rel ZFSet m) : Set (n ⊔ suc m) where
5ed16e2d8eb7 try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 8
diff changeset
34 field
5ed16e2d8eb7 try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 8
diff changeset
35 rel : Rel ZFSet m
5ed16e2d8eb7 try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 8
diff changeset
36 dom : ( y : ZFSet ) → ∀ { x : ZFSet } → rel x y
5ed16e2d8eb7 try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 8
diff changeset
37 fun-eq : { x y z : ZFSet } → ( rel x y ∧ rel x z ) → y ≈ z
5ed16e2d8eb7 try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 8
diff changeset
38
5ed16e2d8eb7 try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 8
diff changeset
39 open Func
5ed16e2d8eb7 try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 8
diff changeset
40
5ed16e2d8eb7 try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 8
diff changeset
41
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
42 record IsZF {n m : Level }
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
43 (ZFSet : Set n)
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
44 (_∋_ : ( A x : ZFSet ) → Set m)
9
5ed16e2d8eb7 try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 8
diff changeset
45 (_≈_ : Rel ZFSet m)
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
46 (∅ : ZFSet)
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
47 (_×_ : ( A B : ZFSet ) → ZFSet)
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
48 (Union : ( A : ZFSet ) → ZFSet)
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
49 (Power : ( A : ZFSet ) → ZFSet)
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
50 (Select : ( ZFSet → Set m ) → ZFSet )
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
51 (Replace : ( ZFSet → ZFSet ) → ZFSet )
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
52 (infinite : ZFSet)
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
53 : Set (suc (n ⊔ m)) where
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
54 field
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
55 isEquivalence : {A B : ZFSet} → IsEquivalence {n} {m} {ZFSet} _≈_
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
56 -- ∀ x ∀ y ∃ z(x ∈ z ∧ y ∈ z)
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
57 pair : ( A B : ZFSet ) → ( (A × B) ∋ A ) ∧ ( (A × B) ∋ B )
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
58 -- ∀ X ∃ A∀ t(t ∈ A ⇔ ∃ x ∈ X(t ∈ x))
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
59 union→ : ( X x y : ZFSet ) → X ∋ x → x ∋ y → Union X ∋ y
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
60 union← : ( X x y : ZFSet ) → Union X ∋ y → X ∋ x → x ∋ y
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
61 _∈_ : ( A B : ZFSet ) → Set m
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
62 A ∈ B = B ∋ A
7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 6
diff changeset
63 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → ∀{ A∋x : Set m } → Set m
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
64 _⊆_ A B {x} {A∋x} = B ∋ x
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
65 _∩_ : ( A B : ZFSet ) → ZFSet
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
66 A ∩ B = Select ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) )
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
67 _∪_ : ( A B : ZFSet ) → ZFSet
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
68 A ∪ B = Select ( λ x → ( A ∋ x ) ∨ ( B ∋ x ) )
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
69 infixr 200 _∈_
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
70 infixr 230 _∩_ _∪_
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
71 infixr 220 _⊆_
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
72 field
4
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 3
diff changeset
73 empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x )
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
74 -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ t ⊆ X ) )
8
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 7
diff changeset
75 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} {y} → _⊆_ t A {x} {y}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 7
diff changeset
76 power← : ∀( A t : ZFSet ) → ∀ {x} {y} → _⊆_ t A {x} {y} → Power A ∋ t
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
77 -- extentionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w )
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
78 extentionality : ( A B z : ZFSet ) → (( A ∋ z ) ⇔ (B ∋ z) ) → A ≈ B
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
79 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) )
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
80 minimul : ZFSet → ZFSet
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
81 regularity : ∀( x : ZFSet ) → ¬ (x ≈ ∅) → ( minimul x ∈ x ∧ ( minimul x ∩ x ≈ ∅ ) )
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
82 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) )
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
83 infinity∅ : ∅ ∈ infinite
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
84 infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ Select ( λ y → x ≈ y )) ∈ infinite
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
85 selection : { ψ : ZFSet → Set m } → ∀ ( y : ZFSet ) → ( y ∈ Select ψ ) → ψ y
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
86 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) )
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
87 replacement : {ψ : ZFSet → ZFSet} → ∀ ( x : ZFSet ) → ( ψ x ∈ Replace ψ )
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
88
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
89 record ZF {n m : Level } : Set (suc (n ⊔ m)) where
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
90 infixr 210 _×_
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
91 infixl 200 _∋_
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
92 infixr 220 _≈_
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
93 field
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
94 ZFSet : Set n
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
95 _∋_ : ( A x : ZFSet ) → Set m
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
96 _≈_ : ( A B : ZFSet ) → Set m
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
97 -- ZF Set constructor
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
98 ∅ : ZFSet
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
99 _×_ : ( A B : ZFSet ) → ZFSet
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
100 Union : ( A : ZFSet ) → ZFSet
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
101 Power : ( A : ZFSet ) → ZFSet
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
102 Select : ( ZFSet → Set m ) → ZFSet
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
103 Replace : ( ZFSet → ZFSet ) → ZFSet
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
104 infinite : ZFSet
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
105 isZF : IsZF ZFSet _∋_ _≈_ ∅ _×_ Union Power Select Replace infinite
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
106
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
107 module zf-exapmle {n m : Level } ( zf : ZF {m} {n} ) where
7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 6
diff changeset
108
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
109 _≈_ = ZF._≈_ zf
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
110 ZFSet = ZF.ZFSet zf
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
111 Select = ZF.Select zf
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
112 ∅ = ZF.∅ zf
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
113 _∩_ = ( IsZF._∩_ ) (ZF.isZF zf)
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
114 _∋_ = ZF._∋_ zf
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
115 replacement = IsZF.replacement ( ZF.isZF zf )
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
116 selection = IsZF.selection ( ZF.isZF zf )
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
117 minimul = IsZF.minimul ( ZF.isZF zf )
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
118 regularity = IsZF.regularity ( ZF.isZF zf )
7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 6
diff changeset
119
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
120 russel : Select ( λ x → x ∋ x ) ≈ ∅
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
121 russel with Select ( λ x → x ∋ x )
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
122 ... | s = {!!}
7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 6
diff changeset
123
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
124 module constructible-set where
7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 6
diff changeset
125
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
126 data Nat : Set zero where
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
127 Zero : Nat
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
128 Suc : Nat → Nat
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
129
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
130 prev : Nat → Nat
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
131 prev (Suc n) = n
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
132 prev Zero = Zero
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
133
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
134 open import Relation.Binary.PropositionalEquality
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
135
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
136 data Transtive {n : Level } : ( lv : Nat) → Set n where
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
137 Φ : {lv : Nat} → Transtive {n} lv
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
138 T-suc : {lv : Nat} → Transtive {n} lv → Transtive lv
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
139 ℵ_ : (lv : Nat) → Transtive (Suc lv)
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
140
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
141 data Constructible {n : Level } {lv : Nat} ( α : Transtive {n} lv ) : Set (suc n) where
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
142 fsub : ( ψ : Transtive {n} lv → Set n ) → Constructible α
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
143 xself : Transtive {n} lv → Constructible α
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
144
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
145 record ConstructibleSet {n : Level } : Set (suc n) where
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
146 field
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
147 level : Nat
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
148 α : Transtive {n} level
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
149 constructible : Constructible α
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
150
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
151 open ConstructibleSet
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
152
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
153 data _c∋_ {n : Level } {lv lv' : Nat} {α : Transtive {n} lv } {α' : Transtive {n} lv' } :
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
154 Constructible {n} {lv} α → Constructible {n} {lv'} α' → Set n where
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
155 xself-fsub : (ta : Transtive {n} lv ) ( ψ : Transtive {n} lv' → Set n ) → (xself ta ) t∋ ( fsub ψ)
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
156 xself-xself : (ta : Transtive {n} lv ) (tx : Transtive {n} lv' ) → (xself ta ) t∋ ( xself tx)
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
157 fsub-fsub : ( ψ : Transtive {n} lv → Set n ) ( ψ₁ : Transtive {n} lv' → Set n ) →( fsub ψ ) t∋ ( fsub ψ₁)
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
158 fsub-xself : ( ψ : Transtive {n} lv → Set n ) (tx : Transtive {n} lv' ) → (fsub ψ ) t∋ ( xself tx)
7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 6
diff changeset
159
10
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
160 _∋_ : {n m : Level} → (ConstructibleSet {n}) → (ConstructibleSet {n} ) → Set m
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
161 _∋_ = {!!}
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
162
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
163
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
164 Transtive→ZF : {n m : Level } → ZF {suc n} {m}
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
165 Transtive→ZF {n} {m} = record {
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
166 ZFSet = ConstructibleSet
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
167 ; _∋_ = _∋_
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
168 ; _≈_ = {!!}
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
169 ; ∅ = record { level = Zero ; α = Φ ; constructible = xself Φ }
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
170 ; _×_ = {!!}
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
171 ; Union = {!!}
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
172 ; Power = {!!}
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
173 ; Select = {!!}
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
174 ; Replace = {!!}
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
175 ; infinite = {!!}
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
176 ; isZF = {!!}
8022e14fce74 add constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 9
diff changeset
177 } where