]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/reverse_complexity/hierarchy.ma
backport of WIP on \lambda\delta to matita 0.99.3
[helm.git] / matita / matita / lib / reverse_complexity / hierarchy.ma
1
2 include "arithmetics/nat.ma".
3 include "arithmetics/log.ma". 
4 include "arithmetics/bigops.ma".
5 include "arithmetics/bounded_quantifiers.ma".
6 include "arithmetics/pidgeon_hole.ma". 
7 include "basics/sets.ma".
8 include "basics/types.ma".
9
10 (************************************ MAX *************************************)
11 notation "Max_{ ident i < n | p } f"
12   with precedence 80
13 for @{'bigop $n max 0 (λ${ident i}. $p) (λ${ident i}. $f)}.
14
15 notation "Max_{ ident i < n } f"
16   with precedence 80
17 for @{'bigop $n max 0 (λ${ident i}.true) (λ${ident i}. $f)}.
18
19 notation "Max_{ ident j ∈ [a,b[ } f"
20   with precedence 80
21 for @{'bigop ($b-$a) max 0 (λ${ident j}.((λ${ident j}.true) (${ident j}+$a)))
22   (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
23   
24 notation "Max_{ ident j ∈ [a,b[ | p } f"
25   with precedence 80
26 for @{'bigop ($b-$a) max 0 (λ${ident j}.((λ${ident j}.$p) (${ident j}+$a)))
27   (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
28
29 lemma Max_assoc: ∀a,b,c. max (max a b) c = max a (max b c).
30 #a #b #c normalize cases (true_or_false (leb a b)) #leab >leab normalize
31   [cases (true_or_false (leb b c )) #lebc >lebc normalize
32     [>(le_to_leb_true a c) // @(transitive_le ? b) @leb_true_to_le //
33     |>leab //
34     ]
35   |cases (true_or_false (leb b c )) #lebc >lebc normalize //
36    >leab normalize >(not_le_to_leb_false a c) // @lt_to_not_le 
37    @(transitive_lt ? b) @not_le_to_lt @leb_false_to_not_le //
38   ]
39 qed.
40
41 lemma Max0 : ∀n. max 0 n = n.
42 // qed.
43
44 lemma Max0r : ∀n. max n 0 = n.
45 #n >commutative_max //
46 qed.
47
48 alias id "max" = "cic:/matita/arithmetics/nat/max#def:2".
49 alias id "mk_Aop" = "cic:/matita/arithmetics/bigops/Aop#con:0:1:2".
50 definition MaxA ≝ 
51   mk_Aop nat 0 max Max0 Max0r (λa,b,c.sym_eq … (Max_assoc a b c)). 
52
53 definition MaxAC ≝ mk_ACop nat 0 MaxA commutative_max.
54
55 lemma le_Max: ∀f,p,n,a. a < n → p a = true →
56   f a ≤  Max_{i < n | p i}(f i).
57 #f #p #n #a #ltan #pa 
58 >(bigop_diff p ? 0 MaxAC f a n) // @(le_maxl … (le_n ?))
59 qed.
60
61 lemma Max_le: ∀f,p,n,b. 
62   (∀a.a < n → p a = true → f a ≤ b) → Max_{i < n | p i}(f i) ≤ b.
63 #f #p #n elim n #b #H // 
64 #b0 #H1 cases (true_or_false (p b)) #Hb
65   [>bigop_Strue [2:@Hb] @to_max [@H1 // | @H #a #ltab #pa @H1 // @le_S //]
66   |>bigop_Sfalse [2:@Hb] @H #a #ltab #pa @H1 // @le_S //
67   ]
68 qed.
69
70 (******************************** big O notation ******************************)
71
72 (*  O f g means g ∈ O(f) *)
73 definition O: relation (nat→nat) ≝
74   λf,g. ∃c.∃n0.∀n. n0 ≤ n → g n ≤ c* (f n).
75   
76 lemma O_refl: ∀s. O s s.
77 #s %{1} %{0} #n #_ >commutative_times <times_n_1 @le_n qed.
78
79 lemma O_trans: ∀s1,s2,s3. O s2 s1 → O s3 s2 → O s3 s1. 
80 #s1 #s2 #s3 * #c1 * #n1 #H1 * #c2 * # n2 #H2 %{(c1*c2)}
81 %{(max n1 n2)} #n #Hmax 
82 @(transitive_le … (H1 ??)) [@(le_maxl … Hmax)]
83 >associative_times @le_times [//|@H2 @(le_maxr … Hmax)]
84 qed.
85
86 lemma sub_O_to_O: ∀s1,s2. O s1 ⊆ O s2 → O s2 s1.
87 #s1 #s2 #H @H // qed.
88
89 lemma O_to_sub_O: ∀s1,s2. O s2 s1 → O s1 ⊆ O s2.
90 #s1 #s2 #H #g #Hg @(O_trans … H) // qed. 
91
92 definition sum_f ≝ λf,g:nat→nat.λn.f n + g n.
93 interpretation "function sum" 'plus f g = (sum_f f g).
94
95 lemma O_plus: ∀f,g,s. O s f → O s g → O s (f+g).
96 #f #g #s * #cf * #nf #Hf * #cg * #ng #Hg
97 %{(cf+cg)} %{(max nf ng)} #n #Hmax normalize 
98 >distributive_times_plus_r @le_plus 
99   [@Hf @(le_maxl … Hmax) |@Hg @(le_maxr … Hmax) ]
100 qed.
101  
102 lemma O_plus_l: ∀f,s1,s2. O s1 f → O (s1+s2) f.
103 #f #s1 #s2 * #c * #a #Os1f %{c} %{a} #n #lean 
104 @(transitive_le … (Os1f n lean)) @le_times //
105 qed.
106
107 lemma O_plus_r: ∀f,s1,s2. O s2 f → O (s1+s2) f.
108 #f #s1 #s2 * #c * #a #Os1f %{c} %{a} #n #lean 
109 @(transitive_le … (Os1f n lean)) @le_times //
110 qed.
111
112 lemma O_absorbl: ∀f,g,s. O s f → O f g → O s (g+f).
113 #f #g #s #Osf #Ofg @(O_plus … Osf) @(O_trans … Osf) //
114 qed.
115
116 lemma O_absorbr: ∀f,g,s. O s f → O f g → O s (f+g).
117 #f #g #s #Osf #Ofg @(O_plus … Osf) @(O_trans … Osf) //
118 qed.
119
120 lemma O_ext2: ∀f,g,s. O s f → (∀x.f x = g x) → O s g.
121 #f #g #s * #c * #a #Osf #eqfg %{c} %{a} #n #lean <eqfg @Osf //
122 qed.    
123
124 definition not_O ≝ λf,g.∀c,n0.∃n. n0 ≤ n ∧ c* (f n) < g n .
125
126 (******************************* small O notation *****************************)
127
128 (*  o f g means g ∈ o(f) *)
129 definition o: relation (nat→nat) ≝
130   λf,g.∀c.∃n0.∀n. n0 ≤ n → c * (g n) < f n.
131   
132 lemma o_irrefl: ∀s. ¬ o s s.
133 #s % #oss cases (oss 1) #n0 #H @(absurd ? (le_n (s n0))) 
134 @lt_to_not_le >(times_n_1 (s n0)) in ⊢ (?%?); >commutative_times @H //
135 qed.
136
137 lemma o_trans: ∀s1,s2,s3. o s2 s1 → o s3 s2 → o s3 s1. 
138 #s1 #s2 #s3 #H1 #H2 #c cases (H1 c) #n1 -H1 #H1 cases (H2 1) #n2 -H2 #H2
139 %{(max n1 n2)} #n #Hmax 
140 @(transitive_lt … (H1 ??)) [@(le_maxl … Hmax)]
141 >(times_n_1 (s2 n)) in ⊢ (?%?); >commutative_times @H2 @(le_maxr … Hmax)
142 qed.
143
144
145 (*********************************** pairing **********************************) 
146
147 axiom pair: nat →nat →nat.
148 axiom fst : nat → nat.
149 axiom snd : nat → nat.
150 axiom fst_pair: ∀a,b. fst (pair a b) = a.
151 axiom snd_pair: ∀a,b. snd (pair a b) = b. 
152
153 interpretation "abstract pair" 'pair f g = (pair f g).
154
155 (************************ basic complexity notions ****************************)
156
157 axiom U: nat → nat → nat → option nat. 
158
159 axiom monotonic_U: ∀i,x,n,m,y.n ≤m →
160   U i x n = Some ? y → U i x m = Some ? y.
161   
162 lemma unique_U: ∀i,x,n,m,yn,ym.
163   U i x n = Some ? yn → U i x m = Some ? ym → yn = ym.
164 #i #x #n #m #yn #ym #Hn #Hm cases (decidable_le n m)
165   [#lenm lapply (monotonic_U … lenm Hn) >Hm #HS destruct (HS) //
166   |#ltmn lapply (monotonic_U … n … Hm) [@lt_to_le @not_le_to_lt //]
167    >Hn #HS destruct (HS) //
168   ]
169 qed.
170
171 definition code_for ≝ λf,i.∀x.
172   ∃n.∀m. n ≤ m → U i x m = f x.
173
174 definition terminate ≝ λi,x,r. ∃y. U i x r = Some ? y.
175 notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. 
176
177 definition lang ≝ λi,x.∃r,y. U i x r = Some ? y ∧ 0  < y. 
178
179 lemma lang_cf :∀f,i,x. code_for f i → 
180   lang i x ↔ ∃y.f x = Some ? y ∧ 0 < y.
181 #f #i #x normalize #H %
182   [* #n * #y * #H1 #posy %{y} % // 
183    cases (H x) -H #m #H <(H (max n m)) [2:@(le_maxr … n) //]
184    @(monotonic_U … H1) @(le_maxl … m) //
185   |cases (H x) -H #m #Hm * #y #Hy %{m} %{y} >Hm // 
186   ]
187 qed.
188
189 (******************************* complexity classes ***************************)
190
191 axiom size: nat → nat.
192 axiom of_size: nat → nat.
193
194 interpretation "size" 'card n = (size n).
195
196 axiom size_of_size: ∀n. |of_size n| = n.
197 axiom monotonic_size: monotonic ? le size.
198
199 axiom of_size_max: ∀i,n. |i| = n → i ≤ of_size n.
200
201 axiom size_fst : ∀n. |fst n| ≤ |n|.
202
203 definition size_f ≝ λf,n.Max_{i < S (of_size n) | eqb (|i|) n}|(f i)|.
204
205 lemma size_f_def: ∀f,n. size_f f n = 
206   Max_{i < S (of_size n) | eqb (|i|) n}|(f i)|.
207 // qed.
208
209 lemma size_f_size : ∀f,n. size_f (f ∘ size) n = |(f n)|.
210 #f #n @le_to_le_to_eq
211   [@Max_le #a #lta #Ha normalize >(eqb_true_to_eq  … Ha) //
212   |<(size_of_size n) in ⊢ (?%?); >size_f_def
213    @(le_Max (λi.|f (|i|)|) ? (S (of_size n)) (of_size n) ??)
214     [@le_S_S // | @eq_to_eqb_true //]
215   ]
216 qed.
217
218 lemma size_f_id : ∀n. size_f (λx.x) n = n.
219 #n @le_to_le_to_eq
220   [@Max_le #a #lta #Ha >(eqb_true_to_eq  … Ha) //
221   |<(size_of_size n) in ⊢ (?%?); >size_f_def
222    @(le_Max (λi.|i|) ? (S (of_size n)) (of_size n) ??)
223     [@le_S_S // | @eq_to_eqb_true //]
224   ]
225 qed.
226
227 lemma size_f_fst : ∀n. size_f fst n ≤ n.
228 #n @Max_le #a #lta #Ha <(eqb_true_to_eq  … Ha) //
229 qed.
230
231 (* C s i means that the complexity of i is O(s) *)
232
233 definition C ≝ λs,i.∃c.∃a.∀x.a ≤ |x| → {i ⊙ x} ↓ (c*(s(|x|))).
234
235 definition CF ≝ λs,f.∃i.code_for f i ∧ C s i.
236
237 lemma ext_CF : ∀f,g,s. (∀n. f n = g n) → CF s f → CF s g.
238 #f #g #s #Hext * #i * #Hcode #HC %{i} %
239   [#x cases (Hcode x) #a #H %{a} <Hext @H | //] 
240 qed. 
241
242 lemma monotonic_CF: ∀s1,s2,f. O s2 s1 → CF s1 f → CF s2 f.
243 #s1 #s2 #f * #c1 * #a #H * #i * #Hcodef #HCs1 %{i} % //
244 cases HCs1 #c2 * #b #H2 %{(c2*c1)} %{(max a b)} 
245 #x #Hmax cases (H2 x ?) [2:@(le_maxr … Hmax)] #y #Hy
246 %{y} @(monotonic_U …Hy) >associative_times @le_times // @H @(le_maxl … Hmax)
247 qed. 
248
249 (*********************** The hierachy theorem (left) **************************)
250
251 theorem hierarchy_theorem_left: ∀s1,s2:nat→nat.
252    O(s1) ⊆ O(s2) → CF s1 ⊆ CF s2.
253 #s1 #s2 #HO #f * #i * #Hcode * #c * #a #Hs1_i %{i} % //
254 cases (sub_O_to_O … HO) -HO #c1 * #b #Hs1s2 
255 %{(c*c1)} %{(max a b)} #x #lemax 
256 cases (Hs1_i x ?) [2: @(le_maxl …lemax)]
257 #y #Hy %{y} @(monotonic_U … Hy) >associative_times
258 @le_times // @Hs1s2 @(le_maxr … lemax)
259 qed.
260
261 (************************** The diagonal language *****************************)
262
263 (* the diagonal language used for the hierarchy theorem *)
264
265 definition diag ≝ λs,i. 
266   U (fst i) i (s (|i|)) = Some ? 0. 
267
268 lemma equiv_diag: ∀s,i. 
269   diag s i ↔ {fst i ⊙ i} ↓ s(|i|) ∧ ¬lang (fst i) i.
270 #s #i %
271   [whd in ⊢ (%→?); #H % [%{0} //] % * #x * #y *
272    #H1 #Hy cut (0 = y) [@(unique_U … H H1)] #eqy /2/
273   |* * #y cases y //
274    #y0 #H * #H1 @False_ind @H1 -H1 whd %{(s (|i|))} %{(S y0)} % //
275   ]
276 qed.
277
278 (* Let us define the characteristic function diag_cf for diag, and prove
279 it correctness *)
280
281 definition diag_cf ≝ λs,i.
282   match U (fst i) i (s (|i|)) with
283   [ None ⇒ None ?
284   | Some y ⇒ if (eqb y 0) then (Some ? 1) else (Some ? 0)].
285
286 lemma diag_cf_OK: ∀s,x. diag s x ↔ ∃y.diag_cf s x = Some ? y ∧ 0 < y.
287 #s #x % 
288   [whd in ⊢ (%→?); #H %{1} % // whd in ⊢ (??%?); >H // 
289   |* #y * whd in ⊢ (??%?→?→%); 
290    cases (U (fst x) x (s (|x|))) normalize
291     [#H destruct
292     |#x cases (true_or_false (eqb x 0)) #Hx >Hx 
293       [>(eqb_true_to_eq … Hx) // 
294       |normalize #H destruct #H @False_ind @(absurd ? H) @lt_to_not_le //  
295       ]
296     ]
297   ]
298 qed.
299
300 lemma diag_spec: ∀s,i. code_for (diag_cf s) i → ∀x. lang i x ↔ diag s x.
301 #s #i #Hcode #x @(iff_trans  … (lang_cf … Hcode)) @iff_sym @diag_cf_OK
302 qed. 
303
304 (******************************************************************************)
305
306 lemma absurd1: ∀P. iff P (¬ P) →False.
307 #P * #H1 #H2 cut (¬P) [% #H2 @(absurd … H2) @H1 //] 
308 #H3 @(absurd ?? H3) @H2 @H3 
309 qed.
310
311 let rec f_img (f:nat →nat) n on n ≝
312   match n with
313   [O ⇒ [ ]
314   |S m ⇒ f m::f_img f m
315   ].
316
317 (* a few lemma to prove injective_to_exists. This is a general result; a nice
318 example of the pidgeon hole pricniple *)
319
320 lemma f_img_to_exists: 
321   ∀f.∀n,a. a ∈ f_img f n → ∃b. b < n ∧ f b = a.
322 #f #n #a elim n normalize [@False_ind]
323 #m #Hind *
324   [#Ha %{m} /2/ |#H cases(Hind H) #b * #Hb #Ha %{b} % // @le_S //]
325 qed. 
326
327 lemma length_f_img: ∀f,n. |f_img f n| = n.
328 #f #n elim n // normalize //
329 qed. 
330
331 lemma unique_f_img: ∀f,n. injective … f → unique ? (f_img f n).
332 #f #n #Hinj elim n normalize //
333 #m #Hind % // % #H lapply(f_img_to_exists …H) * #b * #ltbm
334 #eqbm @(absurd … ltbm) @le_to_not_lt >(Hinj… eqbm) //
335 qed.
336
337 lemma injective_to_exists: ∀f. injective nat nat f → 
338   ∀n.(∀i.i < n → f i < n) → ∀a. a < n → ∃b. b<n ∧ f b = a.
339 #f #finj #n #H1 #a #ltan @(f_img_to_exists f n a)
340 @(eq_length_to_mem_all …  (length_f_img …) (unique_f_img …finj …) …ltan)
341 #x #Hx cases(f_img_to_exists … Hx) #b * #ltbn #eqx <eqx @H1 //
342 qed.
343
344 lemma weak_pad1 :∀n,a.∃b. n ≤ 〈a,b〉. 
345 #n #a 
346 cut (∀i.decidable (〈a,i〉 < n))
347   [#i @decidable_le ] 
348    #Hdec cases(decidable_forall (λb. 〈a,b〉 < n) Hdec n)
349   [#H cut (∀i. i < n → ∃b. b < n ∧ 〈a,b〉 = i)
350     [@(injective_to_exists … H) //]
351    #Hcut %{n} @not_lt_to_le % #Han
352    lapply(Hcut ? Han) * #x * #Hx #Hx2 
353    cut (x = n) [//] #Hxn >Hxn in Hx; /2 by absurd/ 
354   |#H lapply(not_forall_to_exists … Hdec H) 
355    * #b * #H1 #H2 %{b} @not_lt_to_le @H2
356   ]
357 qed.
358
359 lemma pad : ∀n,a. ∃b. n ≤ |〈a,b〉|.
360 #n #a cases (weak_pad1 (of_size n) a) #b #Hb 
361 %{b} <(size_of_size n) @monotonic_size //
362 qed.
363
364 lemma o_to_ex: ∀s1,s2. o s1 s2 → ∀i. C s2 i →
365   ∃b.{i ⊙ 〈i,b〉} ↓ s1 (|〈i,b〉|).
366 #s1 #s2  #H #i * #c * #x0 #H1 
367 cases (H c) #n0 #H2 cases (pad (max x0 n0) i) #b #Hmax
368 %{b} cases (H1 〈i,b〉 ?)
369   [#z #H3 %{z} @(monotonic_U … H3) @lt_to_le @H2
370    @(le_maxr … Hmax)
371   |@(le_maxl … Hmax)
372   ]
373 qed. 
374
375 lemma diag1_not_s1: ∀s1,s2. o s1 s2 → ¬ CF s2 (diag_cf s1).
376 #s1 #s2 #H1 % * #i * #Hcode_i #Hs2_i 
377 cases (o_to_ex  … H1 ? Hs2_i) #b #H2
378 lapply (diag_spec … Hcode_i) #H3
379 @(absurd1 (lang i 〈i,b〉))
380 @(iff_trans … (H3 〈i,b〉)) 
381 @(iff_trans … (equiv_diag …)) >fst_pair 
382 %[* #_ // |#H6 % // ]
383 qed.
384
385 (******************************************************************************)
386
387 definition to_Some ≝ λf.λx:nat. Some nat (f x).
388
389 definition deopt ≝ λn. match n with 
390   [ None ⇒ 1
391   | Some n ⇒ n].
392   
393 definition opt_comp ≝ λf,g:nat → option nat. λx.
394   match g x with 
395   [ None ⇒ None ?
396   | Some y ⇒ f y ].   
397
398 axiom sU2: nat → nat → nat.
399 axiom sU: nat → nat → nat → nat.
400
401 axiom CFU_new: ∀h,g,f,s. 
402   CF s (to_Some h)  → CF s (to_Some g) → CF s (to_Some f) → 
403   O s (λx. sU (size_f h x) (size_f g x) (size_f f x)) → 
404   CF s (λx.U (h x) (g x) (|f x|)).
405     
406 lemma CFU: ∀h,g,f,s1,s2,s3. 
407   CF s1 (to_Some h)  → CF s2 (to_Some g) → CF s3 (to_Some f) → 
408   CF (λx. s1 x + s2 x + s3 x + sU (size_f h x) (size_f g x) (size_f f x)) 
409     (λx.U (h x) (g x) (|f x|)).
410 #h #g #f #s1 #s2 #s3 #Hh #Hg #Hf @CFU_new
411   [@(monotonic_CF … Hh) @O_plus_l @O_plus_l @O_plus_l //
412   |@(monotonic_CF … Hg) @O_plus_l @O_plus_l @O_plus_r //
413   |@(monotonic_CF … Hf) @O_plus_l @O_plus_r //
414   |@O_plus_r //
415   ]
416 qed.
417     
418 axiom monotonic_sU: ∀a1,a2,b1,b2,c1,c2. a1 ≤ a2 → b1 ≤ b2 → c1 ≤c2 →
419   sU a1 b1 c1 ≤ sU a2 b2 c2.
420
421 axiom superlinear_sU: ∀i,x,r. r ≤ sU i x r.
422
423 (* not used *)
424 definition sU_space ≝ λi,x,r.i+x+r.
425 definition sU_time ≝ λi,x,r.i+x+(i^2)*r*(log 2 r). 
426
427 definition IF ≝ λb,f,g:nat →option nat. λx.
428   match b x with 
429   [None ⇒ None ?
430   |Some n ⇒ if (eqb n 0) then f x else g x].
431   
432 axiom IF_CF_new: ∀b,f,g,s. CF s b → CF s f → CF s g → CF s (IF b f g).
433
434 lemma IF_CF: ∀b,f,g,sb,sf,sg. CF sb b → CF sf f → CF sg g → 
435   CF (λn. sb n + sf n + sg n) (IF b f g).
436 #b #f #g #sb #sf #sg #Hb #Hf #Hg @IF_CF_new
437   [@(monotonic_CF … Hb) @O_plus_l @O_plus_l //
438   |@(monotonic_CF … Hf) @O_plus_l @O_plus_r //
439   |@(monotonic_CF … Hg) @O_plus_r //
440   ]
441 qed.
442
443 lemma diag_cf_def : ∀s.∀i. 
444   diag_cf s i =  
445     IF (λi.U (fst i) i (|of_size (s (|i|))|)) (λi.Some ? 1) (λi.Some ? 0) i.
446 #s #i normalize >size_of_size // qed. 
447
448 (* and now ... *)
449 axiom CF_pair: ∀f,g,s. CF s (λx.Some ? (f x)) → CF s (λx.Some ? (g x)) → 
450   CF s (λx.Some ? (pair (f x) (g x))).
451
452 axiom CF_fst: ∀f,s. CF s (λx.Some ? (f x)) → CF s (λx.Some ? (fst (f x))).
453
454 definition sufficiently_large ≝ λs. CF s (λn. Some ? n) ∧ ∀c. CF s (λn. Some ? c).
455
456 definition constructible ≝ λs. CF s (λx.Some ? (of_size (s (|x|)))).
457
458 lemma diag_s: ∀s. sufficiently_large s → constructible s → 
459   CF (λx.sU x x (s x)) (diag_cf s).
460 #s * #Hs_id #Hs_c #Hs_constr 
461 cut (O (λx:ℕ.sU x x (s x)) s) [%{1} %{0} #n //]
462 #O_sU_s @ext_CF [2: #n @sym_eq @diag_cf_def | skip]
463 @IF_CF_new [2,3:@(monotonic_CF … (Hs_c ?)) // ] 
464 @CFU_new
465   [@CF_fst @(monotonic_CF … Hs_id) //
466   |@(monotonic_CF … Hs_id) //
467   |@(monotonic_CF … Hs_constr) //
468   |%{1} %{0} #n #_ >commutative_times <times_n_1
469    @monotonic_sU // >size_f_size >size_of_size //
470   ]
471 qed.