431
|
1 {-# OPTIONS --allow-unsolved-metas #-}
|
|
2 open import Level
|
|
3 open import Relation.Nullary
|
|
4 open import Relation.Binary.PropositionalEquality
|
434
|
5 -- open import Ordinals
|
|
6 module partfunc {n : Level } where -- (O : Ordinals {n}) where
|
431
|
7
|
|
8 open import logic
|
|
9 open import Relation.Binary
|
|
10 open import Data.Empty
|
1175
|
11 open import Data.Unit using ( ⊤ ; tt )
|
431
|
12 open import Data.List hiding (filter)
|
|
13 open import Data.Maybe
|
|
14 open import Relation.Binary
|
|
15 open import Relation.Binary.Core
|
|
16 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ )
|
434
|
17 -- open import filter O
|
431
|
18
|
|
19 open _∧_
|
|
20 open _∨_
|
|
21 open Bool
|
|
22
|
|
23 ----
|
|
24 --
|
|
25 -- Partial Function without ZF
|
|
26 --
|
|
27
|
|
28 record PFunc (Dom : Set n) (Cod : Set n) : Set (suc n) where
|
|
29 field
|
|
30 dom : Dom → Set n
|
|
31 pmap : (x : Dom ) → dom x → Cod
|
|
32 meq : {x : Dom } → { p q : dom x } → pmap x p ≡ pmap x q
|
|
33
|
|
34 ----
|
|
35 --
|
|
36 -- PFunc (Lift n Nat) Cod is equivalent to List (Maybe Cod)
|
|
37 --
|
|
38
|
|
39 data Findp : {Cod : Set n} → List (Maybe Cod) → (x : Nat) → Set where
|
|
40 v0 : {Cod : Set n} → {f : List (Maybe Cod)} → ( v : Cod ) → Findp ( just v ∷ f ) Zero
|
|
41 vn : {Cod : Set n} → {f : List (Maybe Cod)} {d : Maybe Cod} → {x : Nat} → Findp f x → Findp (d ∷ f) (Suc x)
|
|
42
|
|
43 open PFunc
|
|
44
|
|
45 find : {Cod : Set n} → (f : List (Maybe Cod) ) → (x : Nat) → Findp f x → Cod
|
|
46 find (just v ∷ _) 0 (v0 v) = v
|
|
47 find (_ ∷ n) (Suc i) (vn p) = find n i p
|
|
48
|
|
49 findpeq : {Cod : Set n} → (f : List (Maybe Cod)) → {x : Nat} {p q : Findp f x } → find f x p ≡ find f x q
|
|
50 findpeq n {0} {v0 _} {v0 _} = refl
|
|
51 findpeq [] {Suc x} {()}
|
|
52 findpeq (just x₁ ∷ n) {Suc x} {vn p} {vn q} = findpeq n {x} {p} {q}
|
|
53 findpeq (nothing ∷ n) {Suc x} {vn p} {vn q} = findpeq n {x} {p} {q}
|
|
54
|
|
55 List→PFunc : {Cod : Set n} → List (Maybe Cod) → PFunc (Lift n Nat) Cod
|
|
56 List→PFunc fp = record { dom = λ x → Lift n (Findp fp (lower x))
|
|
57 ; pmap = λ x y → find fp (lower x) (lower y)
|
|
58 ; meq = λ {x} {p} {q} → findpeq fp {lower x} {lower p} {lower q}
|
|
59 }
|
|
60 ----
|
|
61 --
|
|
62 -- to List (Maybe Two) is a Latice
|
|
63 --
|
|
64
|
|
65 _3⊆b_ : (f g : List (Maybe Two)) → Bool
|
|
66 [] 3⊆b [] = true
|
|
67 [] 3⊆b (nothing ∷ g) = [] 3⊆b g
|
|
68 [] 3⊆b (_ ∷ g) = true
|
|
69 (nothing ∷ f) 3⊆b [] = f 3⊆b []
|
|
70 (nothing ∷ f) 3⊆b (_ ∷ g) = f 3⊆b g
|
|
71 (just i0 ∷ f) 3⊆b (just i0 ∷ g) = f 3⊆b g
|
|
72 (just i1 ∷ f) 3⊆b (just i1 ∷ g) = f 3⊆b g
|
|
73 _ 3⊆b _ = false
|
|
74
|
|
75 _3⊆_ : (f g : List (Maybe Two)) → Set
|
|
76 f 3⊆ g = (f 3⊆b g) ≡ true
|
|
77
|
|
78 _3∩_ : (f g : List (Maybe Two)) → List (Maybe Two)
|
|
79 [] 3∩ (nothing ∷ g) = nothing ∷ ([] 3∩ g)
|
|
80 [] 3∩ g = []
|
|
81 (nothing ∷ f) 3∩ [] = nothing ∷ f 3∩ []
|
|
82 f 3∩ [] = []
|
|
83 (just i0 ∷ f) 3∩ (just i0 ∷ g) = just i0 ∷ ( f 3∩ g )
|
|
84 (just i1 ∷ f) 3∩ (just i1 ∷ g) = just i1 ∷ ( f 3∩ g )
|
|
85 (_ ∷ f) 3∩ (_ ∷ g) = nothing ∷ ( f 3∩ g )
|
|
86
|
|
87 3∩⊆f : { f g : List (Maybe Two) } → (f 3∩ g ) 3⊆ f
|
|
88 3∩⊆f {[]} {[]} = refl
|
|
89 3∩⊆f {[]} {just _ ∷ g} = refl
|
|
90 3∩⊆f {[]} {nothing ∷ g} = 3∩⊆f {[]} {g}
|
|
91 3∩⊆f {just _ ∷ f} {[]} = refl
|
|
92 3∩⊆f {nothing ∷ f} {[]} = 3∩⊆f {f} {[]}
|
|
93 3∩⊆f {just i0 ∷ f} {just i0 ∷ g} = 3∩⊆f {f} {g}
|
|
94 3∩⊆f {just i1 ∷ f} {just i1 ∷ g} = 3∩⊆f {f} {g}
|
|
95 3∩⊆f {just i0 ∷ f} {just i1 ∷ g} = 3∩⊆f {f} {g}
|
|
96 3∩⊆f {just i1 ∷ f} {just i0 ∷ g} = 3∩⊆f {f} {g}
|
|
97 3∩⊆f {nothing ∷ f} {just _ ∷ g} = 3∩⊆f {f} {g}
|
|
98 3∩⊆f {just i0 ∷ f} {nothing ∷ g} = 3∩⊆f {f} {g}
|
|
99 3∩⊆f {just i1 ∷ f} {nothing ∷ g} = 3∩⊆f {f} {g}
|
|
100 3∩⊆f {nothing ∷ f} {nothing ∷ g} = 3∩⊆f {f} {g}
|
|
101
|
|
102 3∩sym : { f g : List (Maybe Two) } → (f 3∩ g ) ≡ (g 3∩ f )
|
|
103 3∩sym {[]} {[]} = refl
|
|
104 3∩sym {[]} {just _ ∷ g} = refl
|
|
105 3∩sym {[]} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {[]} {g})
|
|
106 3∩sym {just _ ∷ f} {[]} = refl
|
|
107 3∩sym {nothing ∷ f} {[]} = cong (λ k → nothing ∷ k) (3∩sym {f} {[]})
|
|
108 3∩sym {just i0 ∷ f} {just i0 ∷ g} = cong (λ k → just i0 ∷ k) (3∩sym {f} {g})
|
|
109 3∩sym {just i0 ∷ f} {just i1 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g})
|
|
110 3∩sym {just i1 ∷ f} {just i0 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g})
|
|
111 3∩sym {just i1 ∷ f} {just i1 ∷ g} = cong (λ k → just i1 ∷ k) (3∩sym {f} {g})
|
|
112 3∩sym {just i0 ∷ f} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g})
|
|
113 3∩sym {just i1 ∷ f} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g})
|
|
114 3∩sym {nothing ∷ f} {just i0 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g})
|
|
115 3∩sym {nothing ∷ f} {just i1 ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g})
|
|
116 3∩sym {nothing ∷ f} {nothing ∷ g} = cong (λ k → nothing ∷ k) (3∩sym {f} {g})
|
|
117
|
|
118 3⊆-[] : { h : List (Maybe Two) } → [] 3⊆ h
|
|
119 3⊆-[] {[]} = refl
|
|
120 3⊆-[] {just _ ∷ h} = refl
|
|
121 3⊆-[] {nothing ∷ h} = 3⊆-[] {h}
|
|
122
|
|
123 3⊆trans : { f g h : List (Maybe Two) } → f 3⊆ g → g 3⊆ h → f 3⊆ h
|
|
124 3⊆trans {[]} {[]} {[]} f<g g<h = refl
|
|
125 3⊆trans {[]} {[]} {just _ ∷ h} f<g g<h = refl
|
|
126 3⊆trans {[]} {[]} {nothing ∷ h} f<g g<h = 3⊆trans {[]} {[]} {h} refl g<h
|
|
127 3⊆trans {[]} {nothing ∷ g} {[]} f<g g<h = refl
|
|
128 3⊆trans {[]} {just _ ∷ g} {just _ ∷ h} f<g g<h = refl
|
|
129 3⊆trans {[]} {nothing ∷ g} {just _ ∷ h} f<g g<h = refl
|
|
130 3⊆trans {[]} {nothing ∷ g} {nothing ∷ h} f<g g<h = 3⊆trans {[]} {g} {h} f<g g<h
|
|
131 3⊆trans {nothing ∷ f} {[]} {[]} f<g g<h = f<g
|
|
132 3⊆trans {nothing ∷ f} {[]} {just _ ∷ h} f<g g<h = 3⊆trans {f} {[]} {h} f<g (3⊆-[] {h})
|
|
133 3⊆trans {nothing ∷ f} {[]} {nothing ∷ h} f<g g<h = 3⊆trans {f} {[]} {h} f<g g<h
|
|
134 3⊆trans {nothing ∷ f} {nothing ∷ g} {[]} f<g g<h = 3⊆trans {f} {g} {[]} f<g g<h
|
|
135 3⊆trans {nothing ∷ f} {nothing ∷ g} {just _ ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h
|
|
136 3⊆trans {nothing ∷ f} {nothing ∷ g} {nothing ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h
|
|
137 3⊆trans {[]} {just i0 ∷ g} {[]} f<g ()
|
|
138 3⊆trans {[]} {just i1 ∷ g} {[]} f<g ()
|
|
139 3⊆trans {[]} {just x ∷ g} {nothing ∷ h} f<g g<h = 3⊆-[] {h}
|
|
140 3⊆trans {just i0 ∷ f} {[]} {h} () g<h
|
|
141 3⊆trans {just i1 ∷ f} {[]} {h} () g<h
|
|
142 3⊆trans {just x ∷ f} {just i0 ∷ g} {[]} f<g ()
|
|
143 3⊆trans {just x ∷ f} {just i1 ∷ g} {[]} f<g ()
|
|
144 3⊆trans {just i0 ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h
|
|
145 3⊆trans {just i1 ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h
|
|
146 3⊆trans {just x ∷ f} {just i0 ∷ g} {nothing ∷ h} f<g ()
|
|
147 3⊆trans {just x ∷ f} {just i1 ∷ g} {nothing ∷ h} f<g ()
|
|
148 3⊆trans {just i0 ∷ f} {nothing ∷ g} {_} () g<h
|
|
149 3⊆trans {just i1 ∷ f} {nothing ∷ g} {_} () g<h
|
|
150 3⊆trans {nothing ∷ f} {just i0 ∷ g} {[]} f<g ()
|
|
151 3⊆trans {nothing ∷ f} {just i1 ∷ g} {[]} f<g ()
|
|
152 3⊆trans {nothing ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h
|
|
153 3⊆trans {nothing ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g g<h = 3⊆trans {f} {g} {h} f<g g<h
|
|
154 3⊆trans {nothing ∷ f} {just i0 ∷ g} {nothing ∷ h} f<g ()
|
|
155 3⊆trans {nothing ∷ f} {just i1 ∷ g} {nothing ∷ h} f<g ()
|
|
156
|
|
157 3⊆∩f : { f g h : List (Maybe Two) } → f 3⊆ g → f 3⊆ h → f 3⊆ (g 3∩ h )
|
|
158 3⊆∩f {[]} {[]} {[]} f<g f<h = refl
|
|
159 3⊆∩f {[]} {[]} {x ∷ h} f<g f<h = 3⊆-[] {[] 3∩ (x ∷ h)}
|
|
160 3⊆∩f {[]} {x ∷ g} {h} f<g f<h = 3⊆-[] {(x ∷ g) 3∩ h}
|
|
161 3⊆∩f {nothing ∷ f} {[]} {[]} f<g f<h = 3⊆∩f {f} {[]} {[]} f<g f<h
|
|
162 3⊆∩f {nothing ∷ f} {[]} {just _ ∷ h} f<g f<h = f<g
|
|
163 3⊆∩f {nothing ∷ f} {[]} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {[]} {h} f<g f<h
|
|
164 3⊆∩f {just i0 ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h
|
|
165 3⊆∩f {just i1 ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h
|
|
166 3⊆∩f {nothing ∷ f} {just _ ∷ g} {[]} f<g f<h = f<h
|
|
167 3⊆∩f {nothing ∷ f} {just i0 ∷ g} {just i0 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h
|
|
168 3⊆∩f {nothing ∷ f} {just i0 ∷ g} {just i1 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h
|
|
169 3⊆∩f {nothing ∷ f} {just i1 ∷ g} {just i0 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h
|
|
170 3⊆∩f {nothing ∷ f} {just i1 ∷ g} {just i1 ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h
|
|
171 3⊆∩f {nothing ∷ f} {just i0 ∷ g} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h
|
|
172 3⊆∩f {nothing ∷ f} {just i1 ∷ g} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h
|
|
173 3⊆∩f {nothing ∷ f} {nothing ∷ g} {[]} f<g f<h = 3⊆∩f {f} {g} {[]} f<g f<h
|
|
174 3⊆∩f {nothing ∷ f} {nothing ∷ g} {just _ ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h
|
|
175 3⊆∩f {nothing ∷ f} {nothing ∷ g} {nothing ∷ h} f<g f<h = 3⊆∩f {f} {g} {h} f<g f<h
|
|
176
|
|
177 3↑22 : (f : Nat → Two) (i j : Nat) → List (Maybe Two)
|
|
178 3↑22 f Zero j = []
|
|
179 3↑22 f (Suc i) j = just (f j) ∷ 3↑22 f i (Suc j)
|
|
180
|
|
181 _3↑_ : (Nat → Two) → Nat → List (Maybe Two)
|
|
182 _3↑_ f i = 3↑22 f i 0
|
|
183
|
|
184 3↑< : {f : Nat → Two} → { x y : Nat } → x ≤ y → (_3↑_ f x) 3⊆ (_3↑_ f y)
|
|
185 3↑< {f} {x} {y} x<y = lemma x y 0 x<y where
|
|
186 lemma : (x y i : Nat) → x ≤ y → (3↑22 f x i ) 3⊆ (3↑22 f y i )
|
|
187 lemma 0 y i z≤n with f i
|
|
188 lemma Zero Zero i z≤n | i0 = refl
|
|
189 lemma Zero (Suc y) i z≤n | i0 = 3⊆-[] {3↑22 f (Suc y) i}
|
|
190 lemma Zero Zero i z≤n | i1 = refl
|
|
191 lemma Zero (Suc y) i z≤n | i1 = 3⊆-[] {3↑22 f (Suc y) i}
|
|
192 lemma (Suc x) (Suc y) i (s≤s x<y) with f i
|
|
193 lemma (Suc x) (Suc y) i (s≤s x<y) | i0 = lemma x y (Suc i) x<y
|
|
194 lemma (Suc x) (Suc y) i (s≤s x<y) | i1 = lemma x y (Suc i) x<y
|
|
195
|
|
196 Finite3b : (p : List (Maybe Two) ) → Bool
|
|
197 Finite3b [] = true
|
|
198 Finite3b (just _ ∷ f) = Finite3b f
|
|
199 Finite3b (nothing ∷ f) = false
|
|
200
|
|
201 finite3cov : (p : List (Maybe Two) ) → List (Maybe Two)
|
|
202 finite3cov [] = []
|
|
203 finite3cov (just y ∷ x) = just y ∷ finite3cov x
|
|
204 finite3cov (nothing ∷ x) = just i0 ∷ finite3cov x
|
|
205
|
1096
|
206 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
|
|
207 field
|
|
208 filter : L → Set n
|
|
209 f⊆P : PL filter
|
|
210 filter1 : { p q : L } → PL (λ x → q ⊆ x ) → filter p → p ⊆ q → filter q
|
|
211 filter2 : { p q : L } → filter p → filter q → filter (p ∩ q)
|
|
212
|
|
213 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
|
|
214 field
|
|
215 dense : L → Set n
|
|
216 d⊆P : PL dense
|
|
217 dense-f : L → L
|
|
218 dense-d : { p : L} → PL (λ x → p ⊆ x ) → dense ( dense-f p )
|
|
219 dense-p : { p : L} → PL (λ x → p ⊆ x ) → p ⊆ (dense-f p)
|
|
220
|
|
221
|
1175
|
222 Dense-3 : F-Dense (List (Maybe Two) ) (λ x → ⊤ ) _3⊆_ _3∩_
|
431
|
223 Dense-3 = record {
|
|
224 dense = λ x → Finite3b x ≡ true
|
1175
|
225 ; d⊆P = tt
|
431
|
226 ; dense-f = λ x → finite3cov x
|
|
227 ; dense-d = λ {p} d → lemma1 p
|
|
228 ; dense-p = λ {p} d → lemma2 p
|
|
229 } where
|
|
230 lemma1 : (p : List (Maybe Two) ) → Finite3b (finite3cov p) ≡ true
|
|
231 lemma1 [] = refl
|
|
232 lemma1 (just i0 ∷ p) = lemma1 p
|
|
233 lemma1 (just i1 ∷ p) = lemma1 p
|
|
234 lemma1 (nothing ∷ p) = lemma1 p
|
|
235 lemma2 : (p : List (Maybe Two)) → p 3⊆ finite3cov p
|
|
236 lemma2 [] = refl
|
|
237 lemma2 (just i0 ∷ p) = lemma2 p
|
|
238 lemma2 (just i1 ∷ p) = lemma2 p
|
|
239 lemma2 (nothing ∷ p) = lemma2 p
|
|
240
|
|
241 -- min = Data.Nat._⊓_
|
|
242 -- m≤m⊔n = Data.Nat._⊔_
|
|
243 -- open import Data.Nat.Properties
|
|
244
|