From 73428212ec1db9ea1559994f88cd02894a2c9478 Mon Sep 17 00:00:00 2001 From: Andrea Asperti Date: Fri, 3 May 2013 06:51:37 +0000 Subject: [PATCH] reverse complexity --- .../lib/reverse_complexity/hierarchy.ma | 483 ++++++++++++++++++ 1 file changed, 483 insertions(+) create mode 100644 matita/matita/lib/reverse_complexity/hierarchy.ma diff --git a/matita/matita/lib/reverse_complexity/hierarchy.ma b/matita/matita/lib/reverse_complexity/hierarchy.ma new file mode 100644 index 000000000..ff113a0cb --- /dev/null +++ b/matita/matita/lib/reverse_complexity/hierarchy.ma @@ -0,0 +1,483 @@ + +include "arithmetics/nat.ma". +include "arithmetics/log.ma". +(* include "arithmetics/ord.ma". *) +include "arithmetics/bigops.ma". +include "basics/sets.ma". +include "basics/types.ma". + +(************************************ MAX *************************************) +notation "Max_{ ident i < n | p } f" + with precedence 80 +for @{'bigop $n max 0 (λ${ident i}. $p) (λ${ident i}. $f)}. + +notation "Max_{ ident i < n } f" + with precedence 80 +for @{'bigop $n max 0 (λ${ident i}.true) (λ${ident i}. $f)}. + +notation "Max_{ ident j ∈ [a,b[ } f" + with precedence 80 +for @{'bigop ($b-$a) max 0 (λ${ident j}.((λ${ident j}.true) (${ident j}+$a))) + (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}. + +notation "Max_{ ident j ∈ [a,b[ | p } f" + with precedence 80 +for @{'bigop ($b-$a) max 0 (λ${ident j}.((λ${ident j}.$p) (${ident j}+$a))) + (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}. + +lemma Max_assoc: ∀a,b,c. max (max a b) c = max a (max b c). +#a #b #c normalize cases (true_or_false (leb a b)) #leab >leab normalize + [cases (true_or_false (leb b c )) #lebc >lebc normalize + [>(le_to_leb_true a c) // @(transitive_le ? b) @leb_true_to_le // + |>leab // + ] + |cases (true_or_false (leb b c )) #lebc >lebc normalize // + >leab normalize >(not_le_to_leb_false a c) // @lt_to_not_le + @(transitive_lt ? b) @not_le_to_lt @leb_false_to_not_le // + ] +qed. + +lemma Max0 : ∀n. max 0 n = n. +// qed. + +lemma Max0r : ∀n. max n 0 = n. +#n >commutative_max // +qed. + +definition MaxA ≝ + mk_Aop nat 0 max Max0 Max0r (λa,b,c.sym_eq … (Max_assoc a b c)). + +definition MaxAC ≝ mk_ACop nat 0 MaxA commutative_max. + +lemma le_Max: ∀f,p,n,a. a < n → p a = true → + f a ≤ Max_{i < n | p i}(f i). +#f #p #n #a #ltan #pa +>(bigop_diff p ? 0 MaxAC f a n) // @(le_maxl … (le_n ?)) +qed. + +lemma Max_le: ∀f,p,n,b. + (∀a.a < n → p a = true → f a ≤ b) → Max_{i < n | p i}(f i) ≤ b. +#f #p #n elim n #b #H // +#b0 #H1 cases (true_or_false (p b)) #Hb + [>bigop_Strue [2:@Hb] @to_max [@H1 // | @H #a #ltab #pa @H1 // @le_S //] + |>bigop_Sfalse [2:@Hb] @H #a #ltab #pa @H1 // @le_S // + ] +qed. + +(******************************** big O notation ******************************) + +(* O f g means g ∈ O(f) *) +definition O: relation (nat→nat) ≝ + λf,g. ∃c.∃n0.∀n. n0 ≤ n → g n ≤ c* (f n). + +lemma O_refl: ∀s. O s s. +#s %{1} %{0} #n #_ >commutative_times associative_times @le_times [//|@H2 @(le_maxr … Hmax)] +qed. + +lemma sub_O_to_O: ∀s1,s2. O s1 ⊆ O s2 → O s2 s1. +#s1 #s2 #H @H // qed. + +lemma O_to_sub_O: ∀s1,s2. O s2 s1 → O s1 ⊆ O s2. +#s1 #s2 #H #g #Hg @(O_trans … H) // qed. + +definition sum_f ≝ λf,g:nat→nat.λn.f n + g n. +interpretation "function sum" 'plus f g = (sum_f f g). + +lemma O_plus: ∀f,g,s. O s f → O s g → O s (f+g). +#f #g #s * #cf * #nf #Hf * #cg * #ng #Hg +%{(cf+cg)} %{(max nf ng)} #n #Hmax normalize +>distributive_times_plus_r @le_plus + [@Hf @(le_maxl … Hmax) |@Hg @(le_maxr … Hmax) ] +qed. + +lemma O_plus_l: ∀f,s1,s2. O s1 f → O (s1+s2) f. +#f #s1 #s2 * #c * #a #Os1f %{c} %{a} #n #lean +@(transitive_le … (Os1f n lean)) @le_times // +qed. + +lemma O_absorbl: ∀f,g,s. O s f → O f g → O s (g+f). +#f #g #s #Osf #Ofg @(O_plus … Osf) @(O_trans … Osf) // +qed. + +lemma O_absorbr: ∀f,g,s. O s f → O f g → O s (f+g). +#f #g #s #Osf #Ofg @(O_plus … Osf) @(O_trans … Osf) // +qed. + +(* +lemma O_ff: ∀f,s. O s f → O s (f+f). +#f #s #Osf /2/ +qed. *) + +lemma O_ext2: ∀f,g,s. O s f → (∀x.f x = g x) → O s g. +#f #g #s * #c * #a #Osf #eqfg %{c} %{a} #n #lean H normalize #H1 destruct //] +qed. + +lemma Some_to_halt: ∀n,i,y. U i n = Some ? y → u y = None ? . +#n elim n + [#i #y normalize #H destruct (H) + |#m #Hind #i #y normalize + cut (u i = None ? ∨ ∃c. u i = Some ? c) + [cases (u i) [/2/ | #c %2 /2/ ]] + *[#H >H normalize #H1 destruct (H1) // |* #c #H >H normalize @Hind ] + ] +qed. + +lemma monotonici_U: ∀y,n,m,i. + U i m = Some ? y → U i (n+m) = Some ? y. +#y #n #m elim m + [#i normalize #H destruct + |#p #Hind #i H1 normalize // |* #c #H >H normalize #H1 @Hind //] + ] +qed. + +lemma monotonic_U: ∀i,n,m,y.n ≤m → + U i n = Some ? y → U i m = Some ? y. +#i #n #m #y #lenm #H >(plus_minus_m_m m n) // @monotonici_U // +qed. + +(* axiom U: nat → nat → option nat. *) +(* axiom monotonic_U: ∀i,n,m,y.n ≤m → + U i n = Some ? y → U i m = Some ? y. *) + +lemma unique_U: ∀i,n,m,yn,ym. + U i n = Some ? yn → U i m = Some ? ym → yn = ym. +#i #n #m #yn #ym #Hn #Hm cases (decidable_le n m) + [#lenm lapply (monotonic_U … lenm Hn) >Hm #HS destruct (HS) // + |#ltmn lapply (monotonic_U … n … Hm) [@lt_to_le @not_le_to_lt //] + >Hn #HS destruct (HS) // + ] +qed. + +definition code_for ≝ λf,i.∀x. + ∃n.∀m. n ≤ m → U 〈i,x〉 m = f x. + +definition terminate ≝ λc,r. ∃y. U c r = Some ? y. + +interpretation "termination" 'fintersects c r = (terminate c r). + +definition lang ≝ λi,x.∃r,y. U 〈i,x〉 r = Some ? y ∧ 0 < y. + +lemma lang_cf :∀f,i,x. code_for f i → + lang i x ↔ ∃y.f x = Some ? y ∧ 0 < y. +#f #i #x normalize #H % + [* #n * #y * #H1 #posy %{y} % // + cases (H x) -H #m #H <(H (max n m)) [2:@(le_maxr … n) //] + @(monotonic_U … H1) @(le_maxl … m) // + |cases (H x) -H #m #Hm * #y #Hy %{m} %{y} >Hm // + ] +qed. + +(******************************* complexity classes ***************************) + +axiom size: nat → nat. +axiom of_size: nat → nat. + +interpretation "size" 'card n = (size n). + +axiom size_of_size: ∀n. |of_size n| = n. +axiom of_size_max: ∀i,n. |i| = n → i ≤ of_size n. + +axiom size_fst : ∀n. |fst n| ≤ |n|. + +definition size_f ≝ λf,n.Max_{i < S (of_size n) | eqb (|i|) n}|(f i)|. + +lemma size_f_def: ∀f,n. size_f f n = + Max_{i < S (of_size n) | eqb (|i|) n}|(f i)|. +// qed. + +(* +definition Max_const : ∀f,p,n,a. a < n → p a → + ∀n. f n = g n → + Max_{i < n | p n}(f n) = *) + +lemma size_f_size : ∀f,n. size_f (f ∘ size) n = |(f n)|. +#f #n @le_to_le_to_eq + [@Max_le #a #lta #Ha normalize >(eqb_true_to_eq … Ha) // + |<(size_of_size n) in ⊢ (?%?); >size_f_def + @(le_Max (λi.|f (|i|)|) ? (S (of_size n)) (of_size n) ??) + [@le_S_S // | @eq_to_eqb_true //] + ] +qed. + +lemma size_f_id : ∀n. size_f (λx.x) n = n. +#n @le_to_le_to_eq + [@Max_le #a #lta #Ha >(eqb_true_to_eq … Ha) // + |<(size_of_size n) in ⊢ (?%?); >size_f_def + @(le_Max (λi.|i|) ? (S (of_size n)) (of_size n) ??) + [@le_S_S // | @eq_to_eqb_true //] + ] +qed. + +lemma size_f_fst : ∀n. size_f fst n ≤ n. +#n @Max_le #a #lta #Ha <(eqb_true_to_eq … Ha) // +qed. + +(* definition def ≝ λf:nat → option nat.λx.∃y. f x = Some ? y.*) + +(* C s i means that the complexity of i is O(s) *) + +definition C ≝ λs,i.∃c.∃a.∀x.a ≤ |x| → ∃y. + U 〈i,x〉 (c*(s(|x|))) = Some ? y. + +definition CF ≝ λs,f.∃i.code_for f i ∧ C s i. + +lemma ext_CF : ∀f,g,s. (∀n. f n = g n) → CF s f → CF s g. +#f #g #s #Hext * #i * #Hcode #HC %{i} % + [#x cases (Hcode x) #a #H %{a} associative_times @le_times // @H @(le_maxl … Hmax) +qed. + +(************************** The diagonal language *****************************) + +(* the diagonal language used for the hierarchy theorem *) + +definition diag ≝ λs,i. + U 〈fst i,i〉 (s (|i|)) = Some ? 0. + +lemma equiv_diag: ∀s,i. + diag s i ↔ 〈fst i, i〉 ↓ s (|i|) ∧ ¬lang (fst i) i. +#s #i % + [whd in ⊢ (%→?); #H % [%{0} //] % * #x * #y * + #H1 #Hy cut (0 = y) [@(unique_U … H H1)] #eqy /2/ + |* * #y cases y // + #y0 #H * #H1 @False_ind @H1 -H1 whd %{(s (|i|))} %{(S y0)} % // + ] +qed. + +(* Let us define the characteristic function diag_cf for diag, and prove +it correctness *) + +definition diag_cf ≝ λs,i. + match U 〈fst i,i〉 (s (|i|)) with + [ None ⇒ None ? + | Some y ⇒ if (eqb y 0) then (Some ? 1) else (Some ? 0)]. + +lemma diag_cf_OK: ∀s,x. diag s x ↔ ∃y.diag_cf s x = Some ? y ∧ 0 < y. +#s #x % + [whd in ⊢ (%→?); #H %{1} % // whd in ⊢ (??%?); >H // + |* #y * whd in ⊢ (??%?→?→%); + cases (U 〈fst x,x〉 (s (|x|))) normalize + [#H destruct + |#x cases (true_or_false (eqb x 0)) #Hx >Hx + [>(eqb_true_to_eq … Hx) // + |normalize #H destruct #H @False_ind @(absurd ? H) @lt_to_not_le // + ] + ] + ] +qed. + +lemma diag_spec: ∀s,i. code_for (diag_cf s) i → ∀x. lang i x ↔ diag s x. +#s #i #Hcode #x @(iff_trans … (lang_cf … Hcode)) @iff_sym @diag_cf_OK +qed. + +(******************************************************************************) + +lemma absurd1: ∀P. iff P (¬ P) →False. +#P * #H1 #H2 cut (¬P) [% #H2 @(absurd … H2) @H1 //] +#H3 @(absurd ?? H3) @H2 @H3 +qed. + +(* axiom weak_pad : ∀a,∃a0.∀n. a0 < n → ∃b. |〈a,b〉| = n. *) +axiom pad : ∀a,n. |a| < n → ∃b. |〈a,b〉| = n. + +lemma not_included_ex: ∀s1,s2. not_O s2 s1 → ∀i. C s2 i → + ∃b.〈i, 〈i,b〉〉 ↓ s1 (|〈i,b〉|). +#s1 #s2 #H #i * #c * #x0 #H1 +cases (H c (max (S(|i|)) x0)) #x1 * #H2 #H3 cases (pad i x1 ?) + [#b #H4 %{b} + cases (H1 〈i,b〉 ?) + [#z >H4 #H5 %{z} @(monotonic_U … H5) @lt_to_le // + |>H4 @(le_maxr … H2) + ] + |@(le_maxl … H2) + ] +qed. + +lemma diag1_not_s1: ∀s1,s2. not_O s2 s1 → ¬ CF s2 (diag_cf s1). +#s1 #s2 #H1 % * #i * #Hcode_i #Hs2_i +cases (not_included_ex … H1 ? Hs2_i) #b #H2 +lapply (diag_spec … Hcode_i) #H3 +@(absurd1 (lang i 〈i,b〉)) +@(iff_trans … (H3 〈i,b〉)) +@(iff_trans … (equiv_diag …)) >fst_pair +%[* #_ // |#H6 % // ] +qed. + +(******************************************************************************) +(* definition sumF ≝ λf,g:nat→nat.λn.f n + g n. *) + +definition to_Some ≝ λf.λx:nat. Some nat (f x). + +definition deopt ≝ λn. match n with + [ None ⇒ 1 + | Some n ⇒ n]. + +definition opt_comp ≝ λf,g:nat → option nat. λx. + match g x with + [ None ⇒ None ? + | Some y ⇒ f y ]. + +(* axiom CFU: ∀h,g,s. CF s (to_Some h) → CF s (to_Some (of_size ∘ g)) → + CF (Slow s) (λx.U (h x) (g x)). *) + +axiom sU2: nat → nat → nat. +axiom sU: nat → nat → nat → nat. + +(* axiom CFU: CF sU (λx.U (fst x) (snd x)). *) + +axiom CFU: ∀h,g,f,s1,s2,s3. + CF s1 (to_Some h) → CF s2 (to_Some g) → CF s3 (to_Some f) → + CF (λx. s1 x + s2 x + s3 x + sU (size_f h x) (size_f g x) (size_f f x)) + (λx.U 〈h x,g x〉 (|f x|)). + +axiom monotonic_sU: ∀a1,a2,b1,b2,c1,c2. a1 ≤ a2 → b1 ≤ b2 → c1 ≤c2 → + sU a1 b1 c1 ≤ sU a2 b2 c2. + +definition sU_space ≝ λi,x,r.i+x+r. +definition sU_time ≝ λi,x,r.i+x+(i^2)*r*(log 2 r). + +(* +axiom CF_comp: ∀f,g,s1, s2. CF s1 f → CF s2 g → + CF (λx.s2 x + s1 (size (deopt (g x)))) (opt_comp f g). + +(* axiom CF_comp: ∀f,g,s1, s2. CF s1 f → CF s2 g → + CF (s1 ∘ (λx. size (deopt (g x)))) (opt_comp f g). *) + +axiom CF_comp_strong: ∀f,g,s1,s2. CF s1 f → CF s2 g → + CF (s1 ∘ s2) (opt_comp f g). *) + +definition IF ≝ λb,f,g:nat →option nat. λx. + match b x with + [None ⇒ None ? + |Some n ⇒ if (eqb n 0) then f x else g x]. + +axiom IF_CF: ∀b,f,g,sb,sf,sg. CF sb b → CF sf f → CF sg g → + CF (λn. sb n + sf n + sg n) (IF b f g). + +lemma diag_cf_def : ∀s.∀i. + diag_cf s i = + IF (λi.U (pair (fst i) i) (|of_size (s (|i|))|)) (λi.Some ? 1) (λi.Some ? 0) i. +#s #i normalize >size_of_size // qed. + +(* and now ... *) +axiom CF_pair: ∀f,g,s. CF s (λx.Some ? (f x)) → CF s (λx.Some ? (g x)) → + CF s (λx.Some ? (pair (f x) (g x))). + +axiom CF_fst: ∀f,s. CF s (λx.Some ? (f x)) → CF s (λx.Some ? (fst (f x))). + +definition minimal ≝ λs. CF s (λn. Some ? n) ∧ ∀c. CF s (λn. Some ? c). + + +(* +axiom le_snd: ∀n. |snd n| ≤ |n|. +axiom daemon: ∀P:Prop.P. *) + +definition constructible ≝ λs. CF s (λx.Some ? (of_size (s (|x|)))). + +(* +lemma compl1: ∀s. + CF s (to_Some fst) → CF s (to_Some (λx.x)) → CF s (to_Some (λx.(s (|x|)))) → + CF (λx. s x + s x + s x + sU (size_f fst x) (size_f (λx.x) x) (size_f (λx.(s (|x|))) x)) + (λx.U 〈fst x,x〉 (|s (|x|)|)). +#s #H1 #H2 #H3 @CFU // +qed. + +lemma compl1: ∀s. + CF s (to_Some fst) → CF s (to_Some (λx.x)) → CF s (to_Some (λx.(s (|x|)))) → + CF (λx. s x + s x + s x + sU (size_f fst x) (size_f (λx.x) x) (|(s x)| )) + (λx.U 〈fst x,x〉 (|s (|x|)|)). +#s #H1 #H2 #H3 @monotonic_CF [3: @(CFU ??? s s s) @CFU // +qed. *) + +lemma diag_s: ∀s. minimal s → constructible s → + CF (λx.s x + sU x x (s x)) (diag_cf s). +#s * #Hs_id #Hs_c #Hs_constr +@ext_CF [2: #n @sym_eq @diag_cf_def | skip] +@(monotonic_CF ???? (IF_CF (λi:ℕ.U (pair (fst i) i) (|of_size (s (|i|))|)) + … (λi.s i + s i + s i + (sU (size_f fst i) (size_f (λi.i) i) (size_f (λi.of_size (s (|i|))) i))) … (Hs_c 1) (Hs_c 0) … )) + [2: @CFU [@CF_fst // | // |@Hs_constr] + |@(O_ext2 (λn:ℕ.s n+sU (size_f fst n) n (s n) + (s n+s n+s n+s n))) + [2: #i >size_f_size >size_of_size >size_f_id //] + @O_absorbr + [%{1} %{0} #n #_ >commutative_times associative_times +@le_times // @Hs1s2 @(le_maxr … lemax) +qed. + -- 2.39.2