From c51d5f38fa0210a0fb16187f4f39f8de4d296f28 Mon Sep 17 00:00:00 2001 From: Andrea Asperti <andrea.asperti@unibo.it> Date: Tue, 28 Jan 2014 10:09:51 +0000 Subject: [PATCH] progress --- .../lib/reverse_complexity/speed_clean.ma | 272 +++++++++++++----- 1 file changed, 199 insertions(+), 73 deletions(-) diff --git a/matita/matita/lib/reverse_complexity/speed_clean.ma b/matita/matita/lib/reverse_complexity/speed_clean.ma index 470b26695..7c04bfd76 100644 --- a/matita/matita/lib/reverse_complexity/speed_clean.ma +++ b/matita/matita/lib/reverse_complexity/speed_clean.ma @@ -173,6 +173,14 @@ lemma pU_vs_U_None : âi,x,r. pU i x r = â©0,0⪠â U i x r = None ?. |#H >H //] qed. +lemma fst_pU: âi,x,r. fst (pU i x r) = termb i x r. +#i #x #r normalize cases (U i x r) normalize >fst_pair // +qed. + +lemma snd_pU: âi,x,r. snd (pU i x r) = out i x r. +#i #x #r normalize cases (U i x r) normalize >snd_pair // +qed. + (********************************* the speedup ********************************) definition min_input â λh,i,x. μ_{y â [S i,x] } (termb i y (h (S i) y)). @@ -312,8 +320,7 @@ lemma eventually_cancelled: âh,u.¬ânu.âx. nu < x ⧠] qed. - -lemma condition_1_weak: âh,u.g h 0 â g h u. +lemma condition_1: âh,u.g h 0 â g h u. #h #u @(not_to_not ⦠(eventually_cancelled h u)) #H #nu cases (H (max u nu)) #x * #ltx #Hdiff %{x} % [@(le_to_lt_to_lt ⦠ltx) @(le_maxr ⦠(le_n â¦))] @(not_to_not ⦠Hdiff) @@ -352,6 +359,7 @@ space required in addition to dimension of the input. *) axiom MSC : nat â nat. axiom MSC_le: ân. MSC n ⤠n. axiom monotonic_MSC: monotonic ? le MSC. +axiom MSC_pair: âa,b. MSC â©a,b⪠⤠MSC a + MSC b. (* C s i means i is running in O(s) *) @@ -360,7 +368,7 @@ definition C â λs,i.âc.âa.âx.a ⤠x â ây. (* C f s means f â O(s) where MSC âO(s) *) definition CF â λs,f.O s MSC ⧠âi.code_for (total 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 * #HO * #i * #Hcode #HC % // %{i} % [#x cases (Hcode x) #a #H %{a} whd in match (total ??); <Hext @H | //] @@ -413,17 +421,17 @@ lemma CF_comp1: âf,g,s. CF s (total g) â CF s (total f) â (* axiom CF_comp_ext2: âf,g,h,sf,sh. CF sh (total g) â CF sf (total f) â (âx.f(g x) = h x) â - (âx. sf (g x) ⤠sh x) â CF sh (total h). *) + (âx. sf (g x) ⤠sh x) â CF sh (total h). -(* axiom main_MSC: âh,f. CF h f â O h (λx.MSC (f x)). +lemma main_MSC: âh,f. CF h f â O h (λx.MSC (f x)). -axiom CF_S: CF MSC (total S). -axiom CF_fst: CF MSC (total fst). -axiom CF_snd: CF MSC (total snd). +axiom CF_S: CF MSC S. +axiom CF_fst: CF MSC fst. +axiom CF_snd: CF MSC snd. -lemma CF_compS: âh,f. CF h (total f) â CF h (total (S â f)). +lemma CF_compS: âh,f. CF h f â CF h (S â f). #h #f #Hf @(CF_comp ⦠Hf CF_S) @O_plus // @main_MSC // -qed. +qed. lemma CF_comp_fst: âh,f. CF h (total f) â CF h (total (fst â f)). #h #f #Hf @(CF_comp ⦠Hf CF_fst) @O_plus // @main_MSC // @@ -433,9 +441,21 @@ lemma CF_comp_snd: âh,f. CF h (total f) â CF h (total (snd â f)). #h #f #Hf @(CF_comp ⦠Hf CF_snd) @O_plus // @main_MSC // qed. *) +definition id â λx:nat.x. + +axiom CF_id: CF MSC id. axiom CF_compS: âh,f. CF h f â CF h (S â f). axiom CF_comp_fst: âh,f. CF h f â CF h (fst â f). -axiom CF_comp_snd: âh,f. CF h f â CF h (snd â f). +axiom CF_comp_snd: âh,f. CF h f â CF h (snd â f). +axiom CF_comp_pair: âh,f,g. CF h f â CF h g â CF h (λx. â©f x,g xâª). + +lemma CF_fst: CF MSC fst. +@(ext_CF (fst â id)) [#n //] @(CF_comp_fst ⦠CF_id) +qed. + +lemma CF_snd: CF MSC snd. +@(ext_CF (snd â id)) [#n //] @(CF_comp_snd ⦠CF_id) +qed. (************************************** eqb ***********************************) (* definition btotal â @@ -464,18 +484,19 @@ axiom CF_max: âa,b.âp:nat âbool.âf,ha,hb,hp,hf,s. axiom CF_mu: âa,b.âf:nat âbool.âsa,sb,sf,s. CF sa a â CF sb b â CF sf f â O s (λx.sa x + sb x + â_{i â[a x ,S(b x)[ }(sf â©i,xâª)) â - CF s (λx.μ_{i â[a x,b x] }(f â©i,xâª)). + CF s (λx.μ_{i â[a x,b x] }(f â©i,xâª)). + +(****************************** constructibility ******************************) + +definition constructible â λs. CF s s. (********************************* simulation *********************************) axiom sU : nat â nat. -definition termb_unary â λx:â.termb (fst x) (fst (snd x)) (snd (snd x)). axiom monotonic_sU: âi1,i2,x1,x2,s1,s2. i1 ⤠i2 â x1 ⤠x2 â s1 ⤠s2 â sU â©i1,â©x1,s1âªâª ⤠sU â©i2,â©x2,s2âªâª. - -axiom sU_le: âi,x,s. s ⤠sU â©i,â©x,sâªâª. - + lemma monotonic_sU_aux : âx1,x2. fst x1 ⤠fst x2 â fst (snd x1) ⤠fst (snd x2) â snd (snd x1) ⤠snd (snd x2) â sU x1 ⤠sU x2. #x1 #x2 cases (surj_pair x1) #a1 * #y #eqx1 >eqx1 -eqx1 cases (surj_pair y) @@ -484,15 +505,31 @@ cases (surj_pair x2) #a2 * #y2 #eqx2 >eqx2 -eqx2 cases (surj_pair y2) #b2 * #c2 #eqy2 >eqy2 -eqy2 >fst_pair >snd_pair >fst_pair >snd_pair >fst_pair >snd_pair >fst_pair >snd_pair @monotonic_sU qed. + +axiom sU_le: âi,x,s. s ⤠sU â©i,â©x,sâªâª. +axiom sU_le_i: âi,x,s. MSC i ⤠sU â©i,â©x,sâªâª. +axiom sU_le_x: âi,x,s. MSC x ⤠sU â©i,â©x,sâªâª. -axiom CF_termb: CF sU (btotal (termb_unary)). +definition pU_unary â λp. pU (fst p) (fst (snd p)) (snd (snd p)). -axiom CF_compb: âf,g,sf,sg,sh. CF sg (total g) â CF sf (btotal f) â - O sh (λx. sg x + sf (g x)) â CF sh (btotal (f â g)). +axiom CF_U : CF sU pU_unary. + +definition termb_unary â λx:â.termb (fst x) (fst (snd x)) (snd (snd x)). +definition out_unary â λx:â.out (fst x) (fst (snd x)) (snd (snd x)). + +lemma CF_termb: CF sU termb_unary. +@(ext_CF (fst â pU_unary)) [2: @CF_comp_fst @CF_U] +#n whd in ⢠(??%?); whd in ⢠(??(?%)?); >fst_pair % +qed. + +lemma CF_out: CF sU out_unary. +@(ext_CF (snd â pU_unary)) [2: @CF_comp_snd @CF_U] +#n whd in ⢠(??%?); whd in ⢠(??(?%)?); >snd_pair % +qed. (* -lemma CF_termb_comp: âf.CF (sU â f) (btotal (termb_unary â f)). -#f @(CF_compb ⦠CF_termb) *) +lemma CF_termb_comp: âf.CF (sU â f) (termb_unary â f). +#f @(CF_comp ⦠CF_termb) *) (******************** complexity of g ********************) @@ -501,8 +538,7 @@ definition auxg â λh,ux. max_{i â[fst ux,snd ux[ | eqb (min_input h i (snd ux)) (snd ux)} (out i (snd ux) (h (S i) (snd ux))). -lemma compl_g1 : âh,s. - CF s (total (auxg h)) â CF s (total (unary_g h)). +lemma compl_g1 : âh,s. CF s (auxg h) â CF s (unary_g h). #h #s #H1 @(CF_compS ? (auxg h) H1) qed. @@ -517,20 +553,20 @@ qed. lemma compl_g2 : âh,s1,s2,s. CF s1 - (btotal (λp:â.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p)))) â + (λp:â.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) â CF s2 - (total (λp:â.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p))))) â + (λp:â.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) â O s (λx.MSC x + â_{i â[fst x ,snd x[ }(s1 â©i,xâª+s2 â©i,xâª)) â - CF s (total (auxg h)). -#h #s1 #s2 #s #Hs1 #Hs2 #HO @(ext_CF (total (aux1g h))) - [#n whd in ⢠(??%%); @eq_f @eq_aux] + CF s (auxg h). +#h #s1 #s2 #s #Hs1 #Hs2 #HO @(ext_CF (aux1g h)) + [#n whd in ⢠(??%%); @eq_aux] @(CF_max ⦠CF_fst CF_snd Hs1 Hs2 â¦) @(O_trans ⦠HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //] qed. lemma compl_g3 : âh,s. - CF s (total (λp:â.min_input h (fst p) (snd (snd p)))) â - CF s (btotal (λp:â.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p)))). + CF s (λp:â.min_input h (fst p) (snd (snd p))) â + CF s (λp:â.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))). #h #s #H @(CF_eqb ⦠H) @(CF_comp ⦠CF_snd CF_snd) @(O_trans ⦠(proj1 ⦠H)) @O_plus // %{1} %{0} #n #_ >commutative_times <times_n_1 @monotonic_MSC // qed. @@ -559,12 +595,11 @@ qed. *) lemma compl_g4 : âh,s1,s. (CF s1 - (btotal - (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p))))) â + (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)))) â (O s (λx.MSC x + â_{i â[S(fst x) ,S(snd (snd x))[ }(s1 â©i,xâª))) â - CF s (total (λp:â.min_input h (fst p) (snd (snd p)))). -#h #s1 #s #Hs1 #HO @(ext_CF (total (min_input_aux h))) - [#n whd in ⢠(??%%); @eq_f @min_input_eq] + CF s (λp:â.min_input h (fst p) (snd (snd p))). +#h #s1 #s #Hs1 #HO @(ext_CF (min_input_aux h)) + [#n whd in ⢠(??%%); @min_input_eq] @(CF_mu ⦠MSC MSC ⦠Hs1) [@CF_compS @CF_fst |@CF_comp_snd @CF_snd @@ -611,7 +646,6 @@ lemma sigma_bound_decr: âh,a,b. (âa1,a2. a1 ⤠a2 â a2 < b â h a2 ⤠] qed. - lemma coroll: âs1:natânat. (ân. monotonic â le (λa:â.s1 â©a,nâª)) â O (λx.(snd (snd x)-fst x)*(s1 â©snd (snd x),xâª)) (λx.â_{i â[S(fst x) ,S(snd (snd x))[ }(s1 â©i,xâª)). @@ -627,32 +661,58 @@ qed. lemma compl_g5 : âh,s1.(ân. monotonic â le (λa:â.s1 â©a,nâª)) â (CF s1 - (btotal - (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p))))) â + (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)))) â CF (λx.MSC x + (snd (snd x)-fst x)*s1 â©snd (snd x),xâª) - (total (λp:â.min_input h (fst p) (snd (snd p)))). + (λp:â.min_input h (fst p) (snd (snd p))). #h #s1 #Hmono #Hs1 @(compl_g4 ⦠Hs1) @O_plus [@O_plus_l // |@O_plus_r @coroll @Hmono] qed. +(* axiom compl_g6: âh. + (* constructible (λx. h (fst x) (snd x)) â *) + (CF (λx. max (MSC x) (sU â©fst (snd x),â©fst x,h (S (fst (snd x))) (fst x)âªâª)) + (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)))). +*) + +lemma compl_g6: âh. + (âx.MSC xâ¤h (S (fst (snd x))) (fst x)) â + constructible (λx. h (fst x) (snd x)) â (CF (λx. sU â©fst (snd x),â©fst x,h (S (fst (snd x))) (fst x)âªâª) - (btotal - (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p))))). -(* #h #s1 @(ext_CF (btotal (termb_aux h))) + (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)))). +#h #hle #hconstr @(ext_CF (termb_aux h)) [#n normalize >fst_pair >snd_pair >fst_pair >snd_pair // ] -@(CF_compb ⦠CF_termb) *) - +@(CF_comp ⦠(λx.h (S (fst (snd x))) (fst x)) ⦠CF_termb) + [@CF_comp_pair + [@CF_comp_fst @(monotonic_CF ⦠CF_snd) #x // + |@CF_comp_pair + [@(monotonic_CF ⦠CF_fst) #x // + |@(ext_CF ((λx.h (fst x) (snd x))â(λx.â©S (fst (snd x)),fst xâª))) + [#n normalize >fst_pair >snd_pair %] + @(CF_comp ⦠MSC â¦hconstr) + [@CF_comp_pair [@CF_compS @CF_comp_fst // |//] + |@O_plus @le_to_O [//|#n >fst_pair >snd_pair //] + ] + ] + ] + |@O_plus [@le_to_O #n @sU_le | // ] + ] +qed. + + definition faux1 â λh. (λx.MSC x + (snd (snd x)-fst x)*(λx.sU â©fst (snd x),â©fst x,h (S (fst (snd x))) (fst x)âªâª) â©snd (snd x),xâª). (* complexity of min_input *) -lemma compl_g7: âh. (ân. monotonic ? le (h n)) â +lemma compl_g7: âh. + (âx.MSC xâ¤h (S (fst (snd x))) (fst x)) â + constructible (λx. h (fst x) (snd x)) â + (ân. monotonic ? le (h n)) â CF (λx.MSC x + (snd (snd x)-fst x)*sU â©fst x,â©snd (snd x),h (S (fst x)) (snd (snd x))âªâª) - (total (λp:â.min_input h (fst p) (snd (snd p)))). -#h #hmono @(monotonic_CF ⦠(faux1 h)) + (λp:â.min_input h (fst p) (snd (snd p))). +#h #hle #hcostr #hmono @(monotonic_CF ⦠(faux1 h)) [#n normalize >fst_pair >snd_pair //] -@compl_g5 [2:@compl_g6] #n #x #y #lexy >fst_pair >snd_pair +@compl_g5 [2:@(compl_g6 h hle hcostr)] #n #x #y #lexy >fst_pair >snd_pair >fst_pair >snd_pair @monotonic_sU // @hmono @lexy qed. @@ -668,10 +728,13 @@ lemma le_big : âx. x ⤠big x. qed. (* proviamo con x *) -lemma compl_g71: âh. (ân. monotonic ? le (h n)) â +lemma compl_g71: âh. + (âx.MSC xâ¤h (S (fst (snd x))) (fst x)) â + constructible (λx. h (fst x) (snd x)) â + (ân. monotonic ? le (h n)) â CF (λx.MSC (big x) + (snd (snd x)-fst x)*sU â©max (fst x) (snd x),â©snd (snd x),h (S (fst x)) (snd (snd x))âªâª) - (total (λp:â.min_input h (fst p) (snd (snd p)))). -#h #hmono @(monotonic_CF ⦠(compl_g7 h hmono)) #x + (λp:â.min_input h (fst p) (snd (snd p))). +#h #hle #hcostr #hmono @(monotonic_CF ⦠(compl_g7 h hle hcostr hmono)) #x @le_plus [@monotonic_MSC //] cases (decidable_le (fst x) (snd(snd x))) [#Hle @le_times // @monotonic_sU // @(le_maxl ⦠(le_n ⦠)) @@ -679,31 +742,81 @@ cases (decidable_le (fst x) (snd(snd x))) ] qed. +(* axiom compl_g8: âh. CF (λx. sU â©fst x,â©snd (snd x),h (S (fst x)) (snd (snd x))âªâª) - (total (λp:â.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p))))). + (λp:â.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))). *) + +definition out_aux â λh. + out_unary â λp.â©fst p,â©snd(snd p),h (S (fst p)) (snd (snd p))âªâª. + +lemma compl_g8: âh. + constructible (λx. h (fst x) (snd x)) â + (CF (λx. sU â©max (fst x) (snd x),â©snd(snd x),h (S (fst x)) (snd(snd x))âªâª) + (λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p))))). +#h #hconstr @(ext_CF (out_aux h)) + [#n normalize >fst_pair >snd_pair >fst_pair >snd_pair // ] +@(CF_comp ⦠(λx.h (S (fst x)) (snd(snd x)) + MSC x) ⦠CF_out) + [@CF_comp_pair + [@(monotonic_CF ⦠CF_fst) #x // + |@CF_comp_pair + [@CF_comp_snd @(monotonic_CF ⦠CF_snd) #x // + |@(ext_CF ((λx.h (fst x) (snd x))â(λx.â©S (fst x),snd(snd x)âª))) + [#n normalize >fst_pair >snd_pair %] + @(CF_comp ⦠MSC â¦hconstr) + [@CF_comp_pair [@CF_compS // | @CF_comp_snd // ] + |@O_plus @le_to_O [//|#n >fst_pair >snd_pair //] + ] + ] + ] + |@O_plus + [@O_plus + [@le_to_O #n @sU_le + |@(O_trans ⦠(λx.MSC (max (fst x) (snd x)))) + [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx + >fst_pair >snd_pair @(transitive_le ⦠(MSC_pair â¦)) + whd in ⢠(??%); @le_plus + [@monotonic_MSC @(le_maxl ⦠(le_n â¦)) + |>commutative_times <times_n_1 @monotonic_MSC @(le_maxr ⦠(le_n â¦)) + ] + |@le_to_O #x @(transitive_le ???? (sU_le_i ⦠)) // + ] + ] + |@le_to_O #x @monotonic_sU [@(le_maxl ⦠(le_n â¦))|//|//] + ] +qed. +(* lemma compl_g81: âh. + (âx.MSC xâ¤h (S (fst x)) (snd(snd x))) â + constructible (λx. h (fst x) (snd x)) â CF (λx. sU â©max (fst x) (snd x),â©snd (snd x),h (S (fst x)) (snd (snd x))âªâª) - (total (λp:â.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p))))). -#h @(monotonic_CF ⦠(compl_g8 h)) #x @monotonic_sU // @(le_maxl ⦠(le_n ⦠)) -qed. + (λp:â.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))). +#h #hle #hconstr @(monotonic_CF ???? (compl_g8 h hle hconstr)) #x @monotonic_sU // @(le_maxl ⦠(le_n ⦠)) +qed. *) -axiom daemon : False. +(* axiom daemon : False. *) lemma compl_g9 : âh. + (âx.MSC xâ¤h (S (fst (snd x))) (fst x)) â + (âx.MSC xâ¤h (S (fst x)) (snd(snd x))) â + constructible (λx. h (fst x) (snd x)) â (ân. monotonic ? le (h n)) â (ân,a,b. a ⤠b â b ⤠n â h b n ⤠h a n) â CF (λx. (S (snd x-fst x))*MSC â©x,x⪠+ (snd x-fst x)*(S(snd x-fst x))*sU â©x,â©snd x,h (S (fst x)) (snd x)âªâª) - (total (auxg h)). -#h #hmono #hantimono @(compl_g2 h ??? (compl_g3 ⦠(compl_g71 h hmono)) (compl_g81 h)) -@O_plus [@O_plus_l @le_to_O #x elim daemon] + (auxg h). +#h #hle #hle1 #hconstr #hmono #hantimono +@(compl_g2 h ??? (compl_g3 ⦠(compl_g71 h hle hconstr hmono)) (compl_g81 h hle1 hconstr)) +@O_plus + [@O_plus_l @le_to_O #x >(times_n_1 (MSC x)) >commutative_times @le_times + [// | @monotonic_MSC // ]] @(O_trans ⦠(coroll2 ??)) - [#n #a #b #leab #ltb >fst_pair >fst_pair >snd_pair >snd_pair + [#n #a #b #leab #ltb >fst_pair >fst_pair >snd_pair >snd_pair + cut (b ⤠n) [@(transitive_le ⦠(le_snd â¦)) @lt_to_le //] #lebn cut (max a n = n) - [normalize >le_to_leb_true [//|elim daemon (*@(transitive_le ⦠leab lebn)*)]] #maxa - cut (max b n = n) [elim daemon (*normalize >le_to_leb_true //*)] #maxb + [normalize >le_to_leb_true [//|@(transitive_le ⦠leab lebn)]] #maxa + cut (max b n = n) [normalize >le_to_leb_true //] #maxb @le_plus [@le_plus [>big_def >big_def >maxa >maxb //] @le_times @@ -728,11 +841,14 @@ lemma sg_def : âh,a,b. #h #a #b whd in ⢠(??%?); >fst_pair >snd_pair // qed. -lemma compl_g11 : âh. +lemma compl_g11 : âh. + (âx.MSC xâ¤h (S (fst (snd x))) (fst x)) â + (âx.MSC xâ¤h (S (fst x)) (snd(snd x))) â + constructible (λx. h (fst x) (snd x)) â (ân. monotonic ? le (h n)) â (ân,a,b. a ⤠b â b ⤠n â h b n ⤠h a n) â - CF (sg h) (total (unary_g h)). -#h #Hm #Ham @compl_g1 @(compl_g9 h Hm Ham) + CF (sg h) (unary_g h). +#h #hle #hle1 #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hle hle1 hconstr Hm Ham) qed. (**************************** closing the argument ****************************) @@ -768,6 +884,13 @@ lemma h_of_def: âr,a,b.h_of r â©a,b⪠= #r #a #b normalize >fst_pair >snd_pair // qed. +lemma h_le1 : âr.(âx. x ⤠r x) â monotonic ? le r â +(âx:â.MSC xâ¤r (h_of r â©S (fst x),snd (snd x)âª)). +#r #Hr #Hmono #x @(transitive_le ???? (Hr â¦)) +>h_of_def + +(* (âx.MSC xâ¤h (S (fst x)) (snd(snd x))) â *) + lemma mono_h_of_aux: âr.(âx. x ⤠r x) â monotonic ? le r â âd,d1,c,c1,b,b1.c ⤠c1 â d ⤠d1 â b ⤠b1 â h_of_aux r c d b ⤠h_of_aux r c1 d1 b1. @@ -802,7 +925,7 @@ qed. lemma speed_compl: âr:nat ânat. (âx. x ⤠r x) â monotonic ? le r â - CF (h_of r) (total (unary_g (λi,x. r(h_of r â©i,xâª)))). + CF (h_of r) (unary_g (λi,x. r(h_of r â©i,xâª))). #r #Hr #Hmono @(monotonic_CF ⦠(compl_g11 â¦)) [#x cases (surj_pair x) #a * #b #eqx >eqx >sg_def cases (decidable_le b a) @@ -821,7 +944,9 @@ lemma speed_compl: âr:nat ânat. cut (max (S a) b = b) [whd in ⢠(??%?); >(le_to_leb_true ⦠) [%] @not_le_to_lt @ltab] #Hmax1 >Hmax1 - cut (âd.b - a = S d) [elim daemon] * #d #eqd >eqd + cut (âd.b - a = S d) + [%{(pred(b-a))} >S_pred [//] @lt_plus_to_minus_r @not_le_to_lt @ltab] + * #d #eqd >eqd cut (b-S a = d) [//] #eqd1 >eqd1 >h_of_aux_S >eqd1 cut (b - S d = a) [@plus_to_minus >commutative_plus @minus_to_plus @@ -838,25 +963,26 @@ lemma speed_compl: âr:nat ânat. ] qed. +(* lemma unary_g_def : âh,i,x. g h i x = unary_g h â©i,xâª. #h #i #x whd in ⢠(???%); >fst_pair >snd_pair % -qed. +qed. *) (* smn *) axiom smn: âf,s. CF s f â âx. CF (λy.s â©x,yâª) (λy.f â©x,yâª). lemma speed_compl_i: âr:nat ânat. (âx. x ⤠r x) â monotonic ? le r â - âi. CF (λx.h_of r â©i,xâª) (total (λx.g (λi,x. r(h_of r â©i,xâª)) i x)). + âi. CF (λx.h_of r â©i,xâª) (λx.g (λi,x. r(h_of r â©i,xâª)) i x). #r #Hr #Hmono #i -@(ext_CF (total (λx.unary_g (λi,x. r(h_of r â©i,xâª)) â©i,xâª))) - [#n whd in ⢠(??%%); @eq_f @sym_eq @unary_g_def] +@(ext_CF (λx.unary_g (λi,x. r(h_of r â©i,xâª)) â©i,xâª)) + [#n whd in ⢠(??%%); @eq_f @sym_eq >fst_pair >snd_pair %] @smn @(ext_CF ⦠(speed_compl r Hr Hmono)) #n // qed. theorem pseudo_speedup: âr:nat ânat. (âx. x ⤠r x) â monotonic ? le r â - âf.âsf. CF sf (total f) â âg,sg. f â g ⧠CF sg (total g) ⧠O sf (r â sg). + âf.âsf. CF sf f â âg,sg. f â g ⧠CF sg g ⧠O sf (r â sg). (* âm,a.ân. aâ¤n â r(sg a) < m * sf n. *) #r #Hr #Hmono (* f is (g (λi,x. r(h_of r â©i,xâª)) 0) *) @@ -880,7 +1006,7 @@ qed. theorem pseudo_speedup': âr:nat ânat. (âx. x ⤠r x) â monotonic ? le r â - âf.âsf. CF sf (total f) â âg,sg. f â g ⧠CF sg (total g) ⧠+ âf.âsf. CF sf f â âg,sg. f â g ⧠CF sg g ⧠(* ¬ O (r â sg) sf. *) âm,a.ân. aâ¤n â r(sg a) < m * sf n. #r #Hr #Hmono -- 2.39.2