From: Claudio Sacerdoti Coen Date: Fri, 28 Dec 2018 15:25:21 +0000 (+0100) Subject: reverse_complexity lib restored X-Git-Tag: make_still_working~229^2~1^2~7 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=b8e8c61042dd7d4d8bc00971e1ebcd6858064682 reverse_complexity lib restored the toolkit.ma file contains errors. I have commented out a few places --- diff --git a/matita/matita/broken_lib/reverse_complexity/big_O.ma b/matita/matita/broken_lib/reverse_complexity/big_O.ma deleted file mode 100644 index 833572ea2..000000000 --- a/matita/matita/broken_lib/reverse_complexity/big_O.ma +++ /dev/null @@ -1,89 +0,0 @@ - -include "arithmetics/nat.ma". -include "basics/sets.ma". - -(******************************** 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. - -lemma le_to_O: ∀s1,s2. (∀x.s1 x ≤ s2 x) → O s2 s1. -#s1 #s2 #Hle %{1} %{0} #n #_ 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_plus_r: ∀f,s1,s2. O s2 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_times_c: ∀f,c. O f (λx:ℕ.c*f x). -#f #c %{c} %{0} // -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 (times_n_1 (s n0)) in ⊢ (?%?); >commutative_times @H // -qed. - -lemma o_trans: ∀s1,s2,s3. o s2 s1 → o s3 s2 → o s3 s1. -#s1 #s2 #s3 #H1 #H2 #c cases (H1 c) #n1 -H1 #H1 cases (H2 1) #n2 -H2 #H2 -%{(max n1 n2)} #n #Hmax -@(transitive_lt … (H1 ??)) [@(le_maxl … Hmax)] ->(times_n_1 (s2 n)) in ⊢ (?%?); >commutative_times @H2 @(le_maxr … Hmax) -qed. diff --git a/matita/matita/broken_lib/reverse_complexity/gap.ma b/matita/matita/broken_lib/reverse_complexity/gap.ma deleted file mode 100644 index 2c8b6e938..000000000 --- a/matita/matita/broken_lib/reverse_complexity/gap.ma +++ /dev/null @@ -1,251 +0,0 @@ - -include "arithmetics/minimization.ma". -include "arithmetics/bigops.ma". -include "arithmetics/pidgeon_hole.ma". -include "arithmetics/iteration.ma". - -(************************** notation for miminimization ***********************) - -(* an alternative defintion of minimization -definition Min ≝ λa,f. - \big[min,a]_{i < a | f i} i. *) - -notation "μ_{ ident i < n } p" - with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. - -notation "μ_{ ident i ≤ n } p" - with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. - -notation "μ_{ ident i ∈ [a,b] } p" - with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. - -lemma f_min_true: ∀f,a,b. - (∃i. a ≤ i ∧ i ≤ b ∧ f i = true) → f (μ_{i ∈[a,b]} (f i)) = true. -#f #a #b * #i * * #Hil #Hir #Hfi @(f_min_true … (λx. f x)) Hcut in ⊢ (??%); @lt_min %{i} % // % [@Hil |Hm #HS destruct (HS) // - |#ltmn lapply (monotonic_U … n … Hm) [@lt_to_le @not_le_to_lt //] - >Hn #HS destruct (HS) // - ] -qed. - -definition terminate ≝ λi,x,r. ∃y. U i x r = Some ? y. - -notation "〈i,x〉 ↓ r" with precedence 60 for @{terminate $i $x $r}. - -lemma terminate_dec: ∀i,x,n. 〈i,x〉 ↓ n ∨ ¬ 〈i,x〉 ↓ n. -#i #x #n normalize cases (U i x n) - [%2 % * #y #H destruct|#y %1 %{y} //] -qed. - -definition termb ≝ λi,x,t. - match U i x t with [None ⇒ false |Some y ⇒ true]. - -lemma termb_true_to_term: ∀i,x,t. termb i x t = true → 〈i,x〉 ↓ t. -#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] -qed. - -lemma term_to_termb_true: ∀i,x,t. 〈i,x〉 ↓ t → termb i x t = true. -#i #x #t * #y #H normalize >H // -qed. - -lemma decidable_test : ∀n,x,r,r1. - (∀i. i < n → 〈i,x〉 ↓ r ∨ ¬ 〈i,x〉 ↓ r1) ∨ - (∃i. i < n ∧ (¬ 〈i,x〉 ↓ r ∧ 〈i,x〉 ↓ r1)). -#n #x #r1 #r2 - cut (∀i0.decidable ((〈i0,x〉↓r1) ∨ ¬ 〈i0,x〉 ↓ r2)) - [#j @decidable_or [@terminate_dec |@decidable_not @terminate_dec ]] #Hdec - cases(decidable_forall ? Hdec n) - [#H %1 @H - |#H %2 cases (not_forall_to_exists … Hdec H) #j * #leji #Hj - %{j} % // % - [@(not_to_not … Hj) #H %1 @H - |cases (terminate_dec j x r2) // #H @False_ind cases Hj -Hj #Hj - @Hj %2 @H - ] -qed. - -(**************************** the gap theorem *********************************) -definition gapP ≝ λn,x,g,r. ∀i. i < n → 〈i,x〉 ↓ r ∨ ¬ 〈i,x〉 ↓ g r. - -lemma gapP_def : ∀n,x,g,r. - gapP n x g r = ∀i. i < n → 〈i,x〉 ↓ r ∨ ¬ 〈i,x〉 ↓ g r. -// qed. - -lemma upper_bound_aux: ∀g,b,n,x. (∀x. x ≤ g x) → ∀k. - (∃j.j < k ∧ - (∀i. i < n → 〈i,x〉 ↓ g^j b ∨ ¬ 〈i,x〉 ↓ g^(S j) b)) ∨ - ∃l. |l| = k ∧ unique ? l ∧ ∀i. i ∈ l → i < n ∧ 〈i,x〉 ↓ g^k b . -#g#b #n #x #Hg #k elim k - [%2 %{([])} normalize % [% //|#x @False_ind] - |#k0 * - [* #j * #lej #H %1 %{j} % [@le_S // | @H ] - |* #l * * #Hlen #Hunique #Hterm - cases (decidable_test n x (g^k0 b) (g^(S k0) b)) - [#Hcase %1 %{k0} % [@le_n | @Hcase] - |* #j * #ltjn * #H1 #H2 %2 - %{(j::l)} % - [ % [normalize @eq_f @Hlen] whd % // % #H3 - @(absurd ?? H1) @(proj2 … (Hterm …)) @H3 - |#x * - [#eqxj >eqxj % // - |#Hmemx cases(Hterm … Hmemx) #lexn * #y #HU - % [@lexn] %{y} @(monotonic_U ?????? HU) @Hg - ] - ] - ] - ] - ] -qed. - -lemma upper_bound: ∀g,b,n,x. (∀x. x ≤ g x) → ∃r. - (* b ≤ r ∧ r ≤ g^n b ∧ ∀i. i < n → 〈i,x〉 ↓ r ∨ ¬ 〈i,x〉 ↓ g r. *) - b ≤ r ∧ r ≤ g^n b ∧ gapP n x g r. -#g #b #n #x #Hg -cases (upper_bound_aux g b n x Hg n) - [* #j * #Hj #H %{(g^j b)} % [2: @H] % [@le_iter //] - @monotonic_iter2 // @lt_to_le // - |* #l * * #Hlen #Hunique #Hterm %{(g^n b)} % - [% [@le_iter // |@le_n]] - #i #lein %1 @(proj2 … (Hterm ??)) - @(eq_length_to_mem_all … Hlen Hunique … lein) - #x #memx @(proj1 … (Hterm ??)) // - ] -qed. - -definition gapb ≝ λn,x,g,r. - \big[andb,true]_{i < n} ((termb i x r) ∨ ¬(termb i x (g r))). - -lemma gapb_def : ∀n,x,g,r. gapb n x g r = - \big[andb,true]_{i < n} ((termb i x r) ∨ ¬(termb i x (g r))). -// qed. - -lemma gapb_true_to_gapP : ∀n,x,g,r. - gapb n x g r = true → ∀i. i < n → 〈i,x〉 ↓ r ∨ ¬(〈i,x〉 ↓ (g r)). -#n #x #g #r elim n - [>gapb_def >bigop_Strue // - #H #i #lti0 @False_ind @(absurd … lti0) @le_to_not_lt // - |#m #Hind >gapb_def >bigop_Strue // - #H #i #leSm cases (le_to_or_lt_eq … leSm) - [#lem @Hind [@(andb_true_r … H)|@le_S_S_to_le @lem] - |#eqi >(injective_S … eqi) lapply (andb_true_l … H) -H #H cases (orb_true_l … H) -H - [#H %1 @termb_true_to_term // - |#H %2 % #H1 >(term_to_termb_true … H1) in H; normalize #H destruct - ] - ] - ] -qed. - -lemma gapP_to_gapb_true : ∀n,x,g,r. - (∀i. i < n → 〈i,x〉 ↓ r ∨ ¬(〈i,x〉 ↓ (g r))) → gapb n x g r = true. -#n #x #g #r elim n // -#m #Hind #H >gapb_def >bigop_Strue // @true_to_andb_true - [cases (H m (le_n …)) - [#H2 @orb_true_r1 @term_to_termb_true // - |#H2 @orb_true_r2 @sym_eq @noteq_to_eqnot @sym_not_eq - @(not_to_not … H2) @termb_true_to_term - ] - |@Hind #i0 #lei0 @H @le_S // - ] -qed. - - -(* the gap function *) -let rec gap g n on n ≝ - match n with - [ O ⇒ 1 - | S m ⇒ let b ≝ gap g m in μ_{i ∈ [b,g^n b]} (gapb n n g i) - ]. - -lemma gapS: ∀g,m. - gap g (S m) = - (let b ≝ gap g m in - μ_{i ∈ [b,g^(S m) b]} (gapb (S m) (S m) g i)). -// qed. - -lemma upper_bound_gapb: ∀g,m. (∀x. x ≤ g x) → - ∃r:ℕ.gap g m ≤ r ∧ r ≤ g^(S m) (gap g m) ∧ gapb (S m) (S m) g r = true. -#g #m #leg -lapply (upper_bound g (gap g m) (S m) (S m) leg) * #r * * -#H1 #H2 #H3 %{r} % - [% // |@gapP_to_gapb_true @H3] -qed. - -lemma gapS_true: ∀g,m. (∀x. x ≤g x) → gapb (S m) (S m) g (gap g (S m)) = true. -#g #m #leg @(f_min_true (gapb (S m) (S m) g)) @upper_bound_gapb // -qed. - -theorem gap_theorem: ∀g,i.(∀x. x ≤ g x)→∃k.∀n.k < n → - 〈i,n〉 ↓ (gap g n) ∨ ¬ 〈i,n〉 ↓ (g (gap g n)). -#g #i #leg %{i} * - [#lti0 @False_ind @(absurd ?? (not_le_Sn_O i) ) // - |#m #leim lapply (gapS_true g m leg) #H - @(gapb_true_to_gapP … H) // - ] -qed. - -(* an upper bound *) - -let rec sigma n ≝ - match n with - [ O ⇒ 0 | S m ⇒ n + sigma m ]. - -lemma gap_bound: ∀g. (∀x. x ≤ g x) → (monotonic ? le g) → - ∀n.gap g n ≤ g^(sigma n) 1. -#g #leg #gmono #n elim n - [normalize // - |#m #Hind >gapS @(transitive_le ? (g^(S m) (gap g m))) - [@min_up @upper_bound_gapb // - |@(transitive_le ? (g^(S m) (g^(sigma m) 1))) - [@monotonic_iter // |>iter_iter >commutative_plus @le_n - ] - ] -qed. - -lemma gap_bound2: ∀g. (∀x. x ≤ g x) → (monotonic ? le g) → - ∀n.gap g n ≤ g^(n*n) 1. -#g #leg #gmono #n elim n - [normalize // - |#m #Hind >gapS @(transitive_le ? (g^(S m) (gap g m))) - [@min_up @upper_bound_gapb // - |@(transitive_le ? (g^(S m) (g^(m*m) 1))) - [@monotonic_iter // - |>iter_iter @monotonic_iter2 [@leg | normalize 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. - -alias id "max" = "cic:/matita/arithmetics/nat/max#def:2". -alias id "mk_Aop" = "cic:/matita/arithmetics/bigops/Aop#con:0:1:2". -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_plus_r: ∀f,s1,s2. O s2 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 (times_n_1 (s n0)) in ⊢ (?%?); >commutative_times @H // -qed. - -lemma o_trans: ∀s1,s2,s3. o s2 s1 → o s3 s2 → o s3 s1. -#s1 #s2 #s3 #H1 #H2 #c cases (H1 c) #n1 -H1 #H1 cases (H2 1) #n2 -H2 #H2 -%{(max n1 n2)} #n #Hmax -@(transitive_lt … (H1 ??)) [@(le_maxl … Hmax)] ->(times_n_1 (s2 n)) in ⊢ (?%?); >commutative_times @H2 @(le_maxr … Hmax) -qed. - - -(*********************************** pairing **********************************) - -axiom pair: nat →nat →nat. -axiom fst : nat → nat. -axiom snd : nat → nat. -axiom fst_pair: ∀a,b. fst (pair a b) = a. -axiom snd_pair: ∀a,b. snd (pair a b) = b. - -interpretation "abstract pair" 'pair f g = (pair f g). - -(************************ basic complexity notions ****************************) - -(* u is the deterministic configuration relation of the universal machine (one - step) - -axiom u: nat → option nat. - -let rec U c n on n ≝ - match n with - [ O ⇒ None ? - | S m ⇒ match u c with - [ None ⇒ Some ? c (* halting case *) - | Some c1 ⇒ U c1 m - ] - ]. - -lemma halt_U: ∀i,n,y. u i = None ? → U i n = Some ? y → y = i. -#i #n #y #H cases n - [normalize #H1 destruct |#m normalize >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. *) - -axiom U: nat → nat → nat → option nat. -(* -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. *) - -axiom monotonic_U: ∀i,x,n,m,y.n ≤m → - U i x n = Some ? y → U i x 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,x,n,m,yn,ym. - U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. -#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. -notation "[i,x] ↓ r" with precedence 60 for @{terminate $i $x $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 monotonic_size: monotonic ? le size. - -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. *) -lemma weak_pad1 :∀n,a.∃b. n ≤ 〈a,b〉. -#n #a -cut (∀i.decidable (〈a,i〉 < n)) - [#i @decidable_le ] - #Hdec cases(decidable_forall (λb. 〈a,b〉 < n) Hdec n) - [#H cut (∀i. i < n → ∃b. b < n ∧ 〈a,b〉 = i) - [@(injective_to_exists … H) //] - #Hcut %{n} @not_lt_to_le % #Han - lapply(Hcut ? Han) * #x * #Hx #Hx2 - cut (x = n) [//] #Hxn >Hxn in Hx; /2 by absurd/ - |#H lapply(not_forall_to_exists … Hdec H) - * #b * #H1 #H2 %{b} @not_lt_to_le @H2 - ] -qed. - -lemma pad : ∀n,a. ∃b. n ≤ |〈a,b〉|. -#n #a cases (weak_pad1 (of_size n) a) #b #Hb -%{b} <(size_of_size n) @monotonic_size // -qed. - -lemma o_to_ex: ∀s1,s2. o s1 s2 → ∀i. C s2 i → - ∃b.[i, 〈i,b〉] ↓ s1 (|〈i,b〉|). -#s1 #s2 #H #i * #c * #x0 #H1 -cases (H c) #n0 #H2 cases (pad (max x0 n0) i) #b #Hmax -%{b} cases (H1 〈i,b〉 ?) - [#z #H3 %{z} @(monotonic_U … H3) @lt_to_le @H2 - @(le_maxr … Hmax) - |@(le_maxl … Hmax) - ] -qed. - -lemma diag1_not_s1: ∀s1,s2. o s1 s2 → ¬ CF s2 (diag_cf s1). -#s1 #s2 #H1 % * #i * #Hcode_i #Hs2_i -cases (o_to_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 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_new: ∀h,g,f,s. - CF s (to_Some h) → CF s (to_Some g) → CF s (to_Some f) → - O s (λx. sU (size_f h x) (size_f g x) (size_f f x)) → - CF s (λx.U (h x) (g x) (|f x|)). - -lemma 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|)). -#h #g #f #s1 #s2 #s3 #Hh #Hg #Hf @CFU_new - [@(monotonic_CF … Hh) @O_plus_l @O_plus_l @O_plus_l // - |@(monotonic_CF … Hg) @O_plus_l @O_plus_l @O_plus_r // - |@(monotonic_CF … Hf) @O_plus_l @O_plus_r // - |@O_plus_r // - ] -qed. - -axiom monotonic_sU: ∀a1,a2,b1,b2,c1,c2. a1 ≤ a2 → b1 ≤ b2 → c1 ≤c2 → - sU a1 b1 c1 ≤ sU a2 b2 c2. - -axiom superlinear_sU: ∀i,x,r. r ≤ sU i x r. - -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_new: ∀b,f,g,s. CF s b → CF s f → CF s g → CF s (IF b f g). - -lemma 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). -#b #f #g #sb #sf #sg #Hb #Hf #Hg @IF_CF_new - [@(monotonic_CF … Hb) @O_plus_l @O_plus_l // - |@(monotonic_CF … Hf) @O_plus_l @O_plus_r // - |@(monotonic_CF … Hg) @O_plus_r // - ] -qed. - -lemma diag_cf_def : ∀s.∀i. - diag_cf s i = - IF (λi.U (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 diag_s: ∀s. minimal s → constructible s → - CF (λx.sU x x (s x)) (diag_cf s). -#s * #Hs_id #Hs_c #Hs_constr -cut (O (λx:ℕ.sU x x (s x)) s) [%{1} %{0} #n //] -#O_sU_s @ext_CF [2: #n @sym_eq @diag_cf_def | skip] -@IF_CF_new [2,3:@(monotonic_CF … (Hs_c ?)) // ] -@CFU_new - [@CF_fst @(monotonic_CF … Hs_id) // - |@(monotonic_CF … Hs_id) // - |@(monotonic_CF … Hs_constr) // - |%{1} %{0} #n #_ >commutative_times size_f_size >size_of_size // - ] -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] -@IF_CF_new [2,3:@(monotonic_CF … (Hs_c ?)) @O_plus_l //] -@CFU_new - [@CF_fst @(monotonic_CF … Hs_id) @O_plus_l // - |@(monotonic_CF … Hs_id) @O_plus_l // - |@(monotonic_CF … Hs_constr) @O_plus_l // - |@O_plus_r %{1} %{0} #n #_ >commutative_times size_f_size >size_of_size // - ] -qed. *) - -(* proof with old axioms -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. - diff --git a/matita/matita/broken_lib/reverse_complexity/speed_clean.ma b/matita/matita/broken_lib/reverse_complexity/speed_clean.ma deleted file mode 100644 index bfd3d34b1..000000000 --- a/matita/matita/broken_lib/reverse_complexity/speed_clean.ma +++ /dev/null @@ -1,1068 +0,0 @@ -include "basics/types.ma". -include "arithmetics/minimization.ma". -include "arithmetics/bigops.ma". -include "arithmetics/sigma_pi.ma". -include "arithmetics/bounded_quantifiers.ma". -include "reverse_complexity/big_O.ma". -include "basics/core_notation/napart_2.ma". - -(************************* notation for minimization *****************************) -notation "μ_{ ident i < n } p" - with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. - -notation "μ_{ ident i ≤ n } p" - with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. - -notation "μ_{ ident i ∈ [a,b[ } p" - with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}. - -notation "μ_{ ident i ∈ [a,b] } p" - with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. - -(************************************ 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 le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true → - f a ≤ Max_{i ∈ [m,n[ | p i}(f i). -#f #p #n #m #a #lema #ltan #pa ->(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m)) - [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. - -(********************************** pairing ***********************************) -axiom pair: nat → nat → nat. -axiom fst : nat → nat. -axiom snd : nat → nat. - -interpretation "abstract pair" 'pair f g = (pair f g). - -axiom fst_pair: ∀a,b. fst 〈a,b〉 = a. -axiom snd_pair: ∀a,b. snd 〈a,b〉 = b. -axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉. - -axiom le_fst : ∀p. fst p ≤ p. -axiom le_snd : ∀p. snd p ≤ p. -axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉. - -(************************************* U **************************************) -axiom U: nat → nat →nat → option nat. - -axiom monotonic_U: ∀i,x,n,m,y.n ≤m → - U i x n = Some ? y → U i x m = Some ? y. - -lemma unique_U: ∀i,x,n,m,yn,ym. - U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. -#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. - -notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. - -lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n. -#i #x #n normalize cases (U i x n) - [%2 % * #y #H destruct|#y %1 %{y} //] -qed. - -lemma monotonic_terminate: ∀i,x,n,m. - n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m. -#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) // -qed. - -definition termb ≝ λi,x,t. - match U i x t with [None ⇒ false |Some y ⇒ true]. - -lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t. -#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] -qed. - -lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true. -#i #x #t * #y #H normalize >H // -qed. - -definition out ≝ λi,x,r. - match U i x r with [ None ⇒ 0 | Some z ⇒ z]. - -definition bool_to_nat: bool → nat ≝ - λb. match b with [true ⇒ 1 | false ⇒ 0]. - -coercion bool_to_nat. - -definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉. - -lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y. -#i #x #r #y % normalize - [cases (U i x r) normalize - [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H] - #H1 destruct - |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1] - #H1 // - ] - |#H >H //] -qed. - -lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?. -#i #x #r % normalize - [cases (U i x r) normalize // - #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1] - #H1 destruct - |#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)). - -lemma min_input_def : ∀h,i,x. - min_input h i x = min (x -i) (S i) (λy.termb i y (h (S i) y)). -// qed. - -lemma min_input_i: ∀h,i,x. x ≤ i → min_input h i x = S i. -#h #i #x #lexi >min_input_def -cut (x - i = 0) [@sym_eq /2 by eq_minus_O/] #Hcut // -qed. - -lemma min_input_to_terminate: ∀h,i,x. - min_input h i x = x → {i ⊙ x} ↓ (h (S i) x). -#h #i #x #Hminx -cases (decidable_le (S i) x) #Hix - [cases (true_or_false (termb i x (h (S i) x))) #Hcase - [@termb_true_to_term // - |min_input_def in Hminx; #Hminx >Hminx in ⊢ (%→?); - min_input_i in Hminx; - [#eqix >eqix in Hix; * /2/ | @le_S_S_to_le @not_le_to_lt //] - ] -qed. - -lemma min_input_to_lt: ∀h,i,x. - min_input h i x = x → i < x. -#h #i #x #Hminx cases (decidable_le (S i) x) // -#ltxi @False_ind >min_input_i in Hminx; - [#eqix >eqix in ltxi; * /2/ | @le_S_S_to_le @not_le_to_lt //] -qed. - -lemma le_to_min_input: ∀h,i,x,x1. x ≤ x1 → - min_input h i x = x → min_input h i x1 = x. -#h #i #x #x1 #lex #Hminx @(min_exists … (le_S_S … lex)) - [@(fmin_true … (sym_eq … Hminx)) // - |@(min_input_to_lt … Hminx) - |#j #H1 g_def cut (x-u = 0) [/2 by minus_le_minus_minus_comm/] -#eq0 >eq0 normalize // qed. - -lemma g_lt : ∀h,i,x. min_input h i x = x → - out i x (h (S i) x) < g h 0 x. -#h #i #x #H @le_S_S @(le_MaxI … i) /2 by min_input_to_lt/ -qed. - -(* -axiom ax1: ∀h,i. - (∃y.i < y ∧ (termb i y (h (S i) y)=true)) ∨ - ∀y. i < y → (termb i y (h (S i) y)=false). - -lemma eventually_0: ∀h,u.∃nu.∀x. nu < x → - max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) = 0. -#h #u elim u - [%{0} normalize // - |#u0 * #nu0 #Hind cases (ax1 h u0) - [* #x0 * #leu0x0 #Hx0 %{(max nu0 x0)} - #x #Hx >bigop_Sfalse - [>(minus_n_O u0) @Hind @(le_to_lt_to_lt … Hx) /2 by le_maxl/ - |@not_eq_to_eqb_false % #Hf @(absurd (x ≤ x0)) - [bigop_Sfalse - [>(minus_n_O u0) @Hind @(le_to_lt_to_lt … Hx) @le_maxr // - |@not_eq_to_eqb_false >min_input_def - >(min_not_exists (λy.(termb (u0+0) y (h (S (u0+0)) y)))) - [(bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA) - [>H // @(le_to_lt_to_lt …Hx) /2 by le_maxl/ - |@lt_to_le @(le_to_lt_to_lt …Hx) /2 by le_maxr/ - |// - ] -qed. *) - -lemma max_neq0 : ∀a,b. max a b ≠ 0 → a ≠ 0 ∨ b ≠ 0. -#a #b whd in match (max a b); cases (true_or_false (leb a b)) #Hcase >Hcase - [#H %2 @H | #H %1 @H] -qed. - -definition almost_equal ≝ λf,g:nat → nat. ¬ ∀nu.∃x. nu < x ∧ f x ≠ g x. -interpretation "almost equal" 'napart f g = (almost_equal f g). - -lemma eventually_cancelled: ∀h,u.¬∀nu.∃x. nu < x ∧ - max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) ≠ 0. -#h #u elim u - [normalize % #H cases (H u) #x * #_ * #H1 @H1 // - |#u0 @not_to_not #Hind #nu cases (Hind nu) #x * #ltx - cases (true_or_false (eqb (min_input h (u0+O) x) x)) #Hcase - [>bigop_Strue [2:@Hcase] #Hmax cases (max_neq0 … Hmax) -Hmax - [2: #H %{x} % // bigop_Sfalse - [#H %{x1} % [@transitive_lt //| (le_to_min_input … (eqb_true_to_eq … Hcase)) - [@lt_to_not_eq @ltx1 | @lt_to_le @ltx1] - ] - |>bigop_Sfalse [2:@Hcase] #H %{x} % // (bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA) - [>H // |@lt_to_le @(le_to_lt_to_lt …ltx) /2 by le_maxr/ |//] -qed. - -(******************************** Condition 2 *********************************) -definition total ≝ λf.λx:nat. Some nat (f x). - -lemma exists_to_exists_min: ∀h,i. (∃x. i < x ∧ {i ⊙ x} ↓ h (S i) x) → ∃y. min_input h i y = y. -#h #i * #x * #ltix #Hx %{(min_input h i x)} @min_spec_to_min @found // - [@(f_min_true (λy:ℕ.termb i y (h (S i) y))) %{x} % [% // | @term_to_termb_true //] - |#y #leiy #lty @(lt_min_to_false ????? lty) // - ] -qed. - -lemma condition_2: ∀h,i. code_for (total (g h 0)) i → ¬∃x. itermy >Hr -#H @(absurd ? H) @le_to_not_lt @le_n -qed. - - -(********************** complexity ***********************) - -(* We assume operations have a minimal structural complexity MSC. -For instance, for time complexity, MSC is equal to the size of input. -For space complexity, MSC is typically 0, since we only measure the -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) *) - -definition C ≝ λs,i.∃c.∃a.∀x.a ≤ x → ∃y. - U i x (c*(s x)) = Some ? 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 ??); (H m leam) normalize @eq_f @Hext -qed. *) - -lemma monotonic_CF: ∀s1,s2,f.(∀x. s1 x ≤ s2 x) → CF s1 f → CF s2 f. -#s1 #s2 #f #H * #HO * #i * #Hcode #Hs1 % - [cases HO #c * #a -HO #HO %{c} %{a} #n #lean @(transitive_le … (HO n lean)) - @le_times // - |%{i} % [//] cases Hs1 #c * #a -Hs1 #Hs1 %{c} %{a} #n #lean - cases(Hs1 n lean) #y #Hy %{y} @(monotonic_U …Hy) @le_times // - ] -qed. - -lemma O_to_CF: ∀s1,s2,f.O s2 s1 → CF s1 f → CF s2 f. -#s1 #s2 #f #H * #HO * #i * #Hcode #Hs1 % - [@(O_trans … H) // - |%{i} % [//] cases Hs1 #c * #a -Hs1 #Hs1 - cases H #c1 * #a1 #Ha1 %{(c*c1)} %{(a+a1)} #n #lean - cases(Hs1 n ?) [2:@(transitive_le … lean) //] #y #Hy %{y} @(monotonic_U …Hy) - >associative_times @le_times // @Ha1 @(transitive_le … lean) // - ] -qed. - -lemma timesc_CF: ∀s,f,c.CF (λx.c*s x) f → CF s f. -#s #f #c @O_to_CF @O_times_c -qed. - -(********************************* composition ********************************) -axiom CF_comp: ∀f,g,sf,sg,sh. CF sg g → CF sf f → - O sh (λx. sg x + sf (g x)) → CF sh (f ∘ g). - -lemma CF_comp_ext: ∀f,g,h,sh,sf,sg. CF sg g → CF sf f → - (∀x.f(g x) = h x) → O sh (λx. sg x + sf (g x)) → CF sh h. -#f #g #h #sh #sf #sg #Hg #Hf #Heq #H @(ext_CF (f ∘ g)) - [#n normalize @Heq | @(CF_comp … H) //] -qed. - -(* -lemma CF_comp1: ∀f,g,s. CF s (total g) → CF s (total f) → - CF s (total (f ∘ g)). -#f #g #s #Hg #Hf @(timesc_CF … 2) @(monotonic_CF … (CF_comp … Hg Hf)) -*) - -(* -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). - -lemma main_MSC: ∀h,f. CF h f → O h (λx.MSC (f x)). - -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 f → CF h (S ∘ f). -#h #f #Hf @(CF_comp … Hf CF_S) @O_plus // @main_MSC // -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 // -qed. - -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_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 ≝ - λf.λx:nat. match f x with [true ⇒ Some ? 0 |false ⇒ Some ? 1]. *) - -axiom CF_eqb: ∀h,f,g. - CF h f → CF h g → CF h (λx.eqb (f x) (g x)). - -(* -axiom eqb_compl2: ∀h,f,g. - CF2 h (total2 f) → CF2 h (total2 g) → - CF2 h (btotal2 (λx1,x2.eqb (f x1 x2) (g x1 x2))). - -axiom eqb_min_input_compl:∀h,x. - CF (λi.∑_{y ∈ [S i,S x[ }(h i y)) - (btotal (λi.eqb (min_input h i x) x)). *) -(*********************************** maximum **********************************) - -axiom CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s. - CF ha a → CF hb b → CF hp p → CF hf f → - O s (λx.ha x + hb x + ∑_{i ∈[a x ,b x[ }(hp 〈i,x〉 + hf 〈i,x〉)) → - CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). - -(******************************** minimization ********************************) - -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〉)). - -(****************************** constructibility ******************************) - -definition constructible ≝ λs. CF s s. - -lemma constr_comp : ∀s1,s2. constructible s1 → constructible s2 → - (∀x. x ≤ s2 x) → constructible (s2 ∘ s1). -#s1 #s2 #Hs1 #Hs2 #Hle @(CF_comp … Hs1 Hs2) @O_plus @le_to_O #x [@Hle | //] -qed. - -lemma ext_constr: ∀s1,s2. (∀x.s1 x = s2 x) → - constructible s1 → constructible s2. -#s1 #s2 #Hext #Hs1 @(ext_CF … Hext) @(monotonic_CF … Hs1) #x >Hext // -qed. - -(********************************* simulation *********************************) - -axiom sU : nat → nat. - -axiom monotonic_sU: ∀i1,i2,x1,x2,s1,s2. i1 ≤ i2 → x1 ≤ x2 → s1 ≤ s2 → - sU 〈i1,〈x1,s1〉〉 ≤ sU 〈i2,〈x2,s2〉〉. - -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) -#b1 * #c1 #eqy >eqy -eqy -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〉〉. - -definition pU_unary ≝ λp. pU (fst p) (fst (snd p)) (snd (snd p)). - -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) (termb_unary ∘ f). -#f @(CF_comp … CF_termb) *) - -(******************** complexity of g ********************) - -definition unary_g ≝ λh.λux. g h (fst ux) (snd ux). -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 (auxg h) → CF s (unary_g h). -#h #s #H1 @(CF_compS ? (auxg h) H1) -qed. - -definition aux1g ≝ - λh,ux. max_{i ∈[fst ux,snd ux[ | (λp. eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) 〈i,ux〉} - ((λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) 〈i,ux〉). - -lemma eq_aux : ∀h,x.aux1g h x = auxg h x. -#h #x @same_bigop - [#n #_ >fst_pair >snd_pair // |#n #_ #_ >fst_pair >snd_pair //] -qed. - -lemma compl_g2 : ∀h,s1,s2,s. - CF s1 - (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) → - CF s2 - (λ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 (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 (λ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 min_input_def whd in ⊢ (??%?); >minus_S_S @min_f_g #i #_ #_ -whd in ⊢ (??%%); >fst_pair >snd_pair // -qed. - -definition termb_aux ≝ λh. - termb_unary ∘ λp.〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉. - -(* -lemma termb_aux : ∀h,p. - (λx:ℕ.termb (fst x) (fst (snd x)) (snd (snd x))) - 〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉 = - termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)) . -#h #p normalize >fst_pair >snd_pair >fst_pair >snd_pair // -qed. *) - -lemma compl_g4 : ∀h,s1,s. - (CF s1 - (λ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 (λ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 - |@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //] -(* @(ext_CF (btotal (termb_aux h))) - [#n normalize >fst_pair >snd_pair >fst_pair >snd_pair // ] -@(CF_compb … CF_termb) *) -qed. - -(************************* a couple of technical lemmas ***********************) -lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0. -#a elim a // #n #Hind * - [#H @False_ind /2 by absurd/ | #b normalize #H @Hind @le_S_S_to_le /2/] -qed. - -lemma sigma_bound: ∀h,a,b. monotonic nat le h → - ∑_{i ∈ [a,S b[ }(h i) ≤ (S b-a)*h b. -#h #a #b #H cases (decidable_le a b) - [#leab cut (b = pred (S b - a + a)) - [Hb in match (h b); - generalize in match (S b -a); - #n elim n - [// - |#m #Hind >bigop_Strue [2://] @le_plus - [@H @le_n |@(transitive_le … Hind) @le_times [//] @H //] - ] - |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba - cut (S b -a = 0) [@minus_to_0 //] #Hcut >Hcut // - ] -qed. - -lemma sigma_bound_decr: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) → - ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a. -#h #a #b #H cases (decidable_le a b) - [#leab cut ((b -a) +a ≤ b) [/2 by le_minus_to_plus_r/] generalize in match (b -a); - #n elim n - [// - |#m #Hind >bigop_Strue [2://] #Hm - cut (m+a ≤ b) [@(transitive_le … Hm) //] #Hm1 - @le_plus [@H // |@(transitive_le … (Hind Hm1)) //] - ] - |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba - cut (b -a = 0) [@minus_to_0 @lt_to_le @ltba] #Hcut >Hcut // - ] -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〉)). -#s1 #Hs1 %{1} %{0} #n #_ >commutative_times minus_S_S //] -qed. - -lemma coroll2: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) → -O (λx.(snd x - fst x)*s1 〈fst x,x〉) (λx.∑_{i ∈[fst x,snd x[ }(s1 〈i,x〉)). -#s1 #Hs1 %{1} %{0} #n #_ >commutative_times fst_pair >snd_pair >fst_pair >snd_pair // ] -@(CF_comp … (λx.MSC 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 - [@O_plus - [@(O_trans … (λx.MSC (fst x) + MSC (max (fst (snd x)) (snd (snd x))))) - [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx - >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) - >distributive_times_plus @le_plus [//] - cases (surj_pair b) #c * #d #eqb >eqb - >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) - whd in ⊢ (??%); @le_plus - [@monotonic_MSC @(le_maxl … (le_n …)) - |>commutative_times fst_pair >snd_pair //] -@compl_g5 [2:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair ->fst_pair >snd_pair @monotonic_sU // @hmono @lexy -qed.*) - -definition big : nat →nat ≝ λx. - let m ≝ max (fst x) (snd x) in 〈m,m〉. - -lemma big_def : ∀a,b. big 〈a,b〉 = 〈max a b,max a b〉. -#a #b normalize >fst_pair >snd_pair // qed. - -lemma le_big : ∀x. x ≤ big x. -#x cases (surj_pair x) #a * #b #eqx >eqx @le_pair >fst_pair >snd_pair -[@(le_maxl … (le_n …)) | @(le_maxr … (le_n …))] -qed. - -definition faux2 ≝ λh. - (λx.MSC x + (snd (snd x)-fst x)* - (λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉). - -(* proviamo con x *) -lemma compl_g7: ∀h. - constructible (λx. h (fst x) (snd x)) → - (∀n. monotonic ? le (h n)) → - CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉) - (λp:ℕ.min_input h (fst p) (snd (snd p))). -#h #hcostr #hmono @(monotonic_CF … (faux2 h)) - [#n normalize >fst_pair >snd_pair //] -@compl_g5 [2:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair ->fst_pair >snd_pair @monotonic_sU // @hmono @lexy -qed. - -(* proviamo con x *) -lemma compl_g71: ∀h. - 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))〉〉) - (λp:ℕ.min_input h (fst p) (snd (snd p))). -#h #hcostr #hmono @(monotonic_CF … (compl_g7 h hcostr hmono)) #x -@le_plus [@monotonic_MSC //] -cases (decidable_le (fst x) (snd(snd x))) - [#Hle @le_times // @monotonic_sU - |#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt] - ] -qed. - -(* -axiom compl_g8: ∀h. - CF (λx. sU 〈fst 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)))). *) - -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 (MSC x)) >commutative_times @le_times - [// | @monotonic_MSC // ]] -@(O_trans … (coroll2 ??)) - [#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 [//|@(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 - [/2 by monotonic_le_minus_r/ - |@monotonic_sU // @hantimono [@le_S_S // |@ltb] - ] - |@monotonic_sU // @hantimono [@le_S_S // |@ltb] - ] - |@le_to_O #n >fst_pair >snd_pair - cut (max (fst n) n = n) [normalize >le_to_leb_true //] #Hmax >Hmax - >associative_plus >distributive_times_plus - @le_plus [@le_times [@le_S // |>big_def >Hmax //] |//] - ] -qed. - -definition sg ≝ λh,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)〉〉. - -lemma sg_def : ∀h,a,b. - sg h 〈a,b〉 = (S (b-a))*MSC 〈〈a,b〉,〈a,b〉〉 + - (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉. -#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair // -qed. - -lemma compl_g11 : ∀h. - 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) (unary_g h). -#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham) -qed. - -(**************************** closing the argument ****************************) - -let rec h_of_aux (r:nat →nat) (c,d,b:nat) on d : nat ≝ - match d with - [ O ⇒ c (* MSC 〈〈b,b〉,〈b,b〉〉 *) - | S d1 ⇒ (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + - d*(S d)*sU 〈〈b-d,b〉,〈b,r (h_of_aux r c d1 b)〉〉]. - -lemma h_of_aux_O: ∀r,c,b. - h_of_aux r c O b = c. -// qed. - -lemma h_of_aux_S : ∀r,c,d,b. - h_of_aux r c (S d) b = - (S (S d))*(MSC 〈〈b-(S d),b〉,〈b-(S d),b〉〉) + - (S d)*(S (S d))*sU 〈〈b-(S d),b〉,〈b,r(h_of_aux r c d b)〉〉. -// qed. - -definition h_of ≝ λr,p. - let m ≝ max (fst p) (snd p) in - h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (snd p - fst p) (snd p). - -lemma h_of_O: ∀r,a,b. b ≤ a → - h_of r 〈a,b〉 = let m ≝ max a b in MSC 〈〈m,m〉,〈m,m〉〉. -#r #a #b #Hle normalize >fst_pair >snd_pair >(minus_to_0 … Hle) // -qed. - -lemma h_of_def: ∀r,a,b.h_of r 〈a,b〉 = - let m ≝ max a b in - h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (b - a) b. -#r #a #b normalize >fst_pair >snd_pair // -qed. - -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. -#r #Hr #monor #d #d1 lapply d -d elim d1 - [#d #c #c1 #b #b1 #Hc #Hd @(le_n_O_elim ? Hd) #leb - >h_of_aux_O >h_of_aux_O // - |#m #Hind #d #c #c1 #b #b1 #lec #led #leb cases (le_to_or_lt_eq … led) - [#ltd @(transitive_le … (Hind … lec ? leb)) [@le_S_S_to_le @ltd] - >h_of_aux_S @(transitive_le ???? (le_plus_n …)) - >(times_n_1 (h_of_aux r c1 m b1)) in ⊢ (?%?); - >commutative_times @le_times [//|@(transitive_le … (Hr ?)) @sU_le] - |#Hd >Hd >h_of_aux_S >h_of_aux_S - cut (b-S m ≤ b1 - S m) [/2 by monotonic_le_minus_l/] #Hb1 - @le_plus [@le_times //] - [@monotonic_MSC @le_pair @le_pair // - |@le_times [//] @monotonic_sU - [@le_pair // |// |@monor @Hind //] - ] - ] - ] -qed. - -lemma mono_h_of2: ∀r.(∀x. x ≤ r x) → monotonic ? le r → - ∀i,b,b1. b ≤ b1 → h_of r 〈i,b〉 ≤ h_of r 〈i,b1〉. -#r #Hr #Hmono #i #a #b #leab >h_of_def >h_of_def -cut (max i a ≤ max i b) - [@to_max - [@(le_maxl … (le_n …))|@(transitive_le … leab) @(le_maxr … (le_n …))]] -#Hmax @(mono_h_of_aux r Hr Hmono) - [@monotonic_MSC @le_pair @le_pair @Hmax |/2 by monotonic_le_minus_l/ |@leab] -qed. - -axiom h_of_constr : ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - constructible (h_of r). - -lemma speed_compl: ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))). -#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …)) - [#x cases (surj_pair x) #a * #b #eqx >eqx - >sg_def cases (decidable_le b a) - [#leba >(minus_to_0 … leba) normalize in ⊢ (?%?); - h_of_def - cut (max a b = a) - [normalize cases (le_to_or_lt_eq … leba) - [#ltba >(lt_to_leb_false … ltba) % - |#eqba (le_to_leb_true … (le_n ?)) % ]] - #Hmax >Hmax normalize >(minus_to_0 … leba) normalize - @monotonic_MSC @le_pair @le_pair // - |#ltab >h_of_def >h_of_def - cut (max a b = b) - [normalize >(le_to_leb_true … ) [%] @lt_to_le @not_le_to_lt @ltab] - #Hmax >Hmax - 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) - [%{(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 - [@lt_to_le @not_le_to_lt // | //]] #eqd2 >eqd2 - normalize // - ] - |#n #a #b #leab #lebn >h_of_def >h_of_def - cut (max a n = n) - [normalize >le_to_leb_true [%|@(transitive_le … leab lebn)]] #Hmaxa - cut (max b n = n) - [normalize >(le_to_leb_true … lebn) %] #Hmaxb - >Hmaxa >Hmaxb @Hmono @(mono_h_of_aux r … Hr Hmono) // /2 by monotonic_le_minus_r/ - |#n #a #b #leab @Hmono @(mono_h_of2 … Hr Hmono … leab) - |@(constr_comp … Hconstr Hr) @(ext_constr (h_of r)) - [#x cases (surj_pair x) #a * #b #eqx >eqx >fst_pair >snd_pair //] - @(h_of_constr r Hr Hmono Hconstr) - ] -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. *) - -(* 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 → constructible r → - ∀i. CF (λx.h_of r 〈i,x〉) (λx.g (λi,x. r(h_of r 〈i,x〉)) i x). -#r #Hr #Hmono #Hconstr #i -@(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 Hconstr)) #n // -qed. - -theorem pseudo_speedup: - ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∃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 #Hconstr -(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * -#Hcodei #HCi -(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} -(* sg is (λx.h_of r 〈i,x〉) *) -%{(λx. h_of r 〈S i,x〉)} -lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg -%[%[@condition_1 |@Hg] - |cases Hg #H1 * #j * #Hcodej #HCj - lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) - cases HCi #m * #a #Ha %{m} %{(max (S i) a)} #n #ltin @lt_to_le @not_le_to_lt - @(not_to_not … Hcond2) -Hcond2 #Hlesf %{n} % - [@(transitive_le … ltin) @(le_maxl … (le_n …))] - cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] - #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) // - ] -qed. - -theorem pseudo_speedup': - ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∃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 #Hconstr -(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * -#Hcodei #HCi -(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} -(* sg is (λx.h_of r 〈i,x〉) *) -%{(λx. h_of r 〈S i,x〉)} -lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg -%[%[@condition_1 |@Hg] - |cases Hg #H1 * #j * #Hcodej #HCj - lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) - cases HCi #m * #a #Ha - %{m} %{(max (S i) a)} #n #ltin @not_le_to_lt @(not_to_not … Hcond2) -Hcond2 #Hlesf - %{n} % [@(transitive_le … ltin) @(le_maxl … (le_n …))] - cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] - #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) - @Hmono @(mono_h_of2 … Hr Hmono … ltin) - ] -qed. - diff --git a/matita/matita/broken_lib/reverse_complexity/speed_def.ma b/matita/matita/broken_lib/reverse_complexity/speed_def.ma deleted file mode 100644 index 9812cfb08..000000000 --- a/matita/matita/broken_lib/reverse_complexity/speed_def.ma +++ /dev/null @@ -1,922 +0,0 @@ -include "basics/types.ma". -include "arithmetics/minimization.ma". -include "arithmetics/bigops.ma". -include "arithmetics/sigma_pi.ma". -include "arithmetics/bounded_quantifiers.ma". -include "reverse_complexity/big_O.ma". -include "basics/core_notation/napart_2.ma". - -(************************* notation for minimization *****************************) -notation "μ_{ ident i < n } p" - with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. - -notation "μ_{ ident i ≤ n } p" - with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. - -notation "μ_{ ident i ∈ [a,b[ } p" - with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}. - -notation "μ_{ ident i ∈ [a,b] } p" - with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. - -(************************************ 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 le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true → - f a ≤ Max_{i ∈ [m,n[ | p i}(f i). -#f #p #n #m #a #lema #ltan #pa ->(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m)) - [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. - -(********************************** pairing ***********************************) -axiom pair: nat → nat → nat. -axiom fst : nat → nat. -axiom snd : nat → nat. - -interpretation "abstract pair" 'pair f g = (pair f g). - -axiom fst_pair: ∀a,b. fst 〈a,b〉 = a. -axiom snd_pair: ∀a,b. snd 〈a,b〉 = b. -axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉. - -axiom le_fst : ∀p. fst p ≤ p. -axiom le_snd : ∀p. snd p ≤ p. -axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉. - -(************************************* U **************************************) -axiom U: nat → nat →nat → option nat. - -axiom monotonic_U: ∀i,x,n,m,y.n ≤m → - U i x n = Some ? y → U i x m = Some ? y. - -lemma unique_U: ∀i,x,n,m,yn,ym. - U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. -#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. - -notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. - -lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n. -#i #x #n normalize cases (U i x n) - [%2 % * #y #H destruct|#y %1 %{y} //] -qed. - -lemma monotonic_terminate: ∀i,x,n,m. - n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m. -#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) // -qed. - -definition termb ≝ λi,x,t. - match U i x t with [None ⇒ false |Some y ⇒ true]. - -lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t. -#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] -qed. - -lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true. -#i #x #t * #y #H normalize >H // -qed. - -definition out ≝ λi,x,r. - match U i x r with [ None ⇒ 0 | Some z ⇒ z]. - -definition bool_to_nat: bool → nat ≝ - λb. match b with [true ⇒ 1 | false ⇒ 0]. - -coercion bool_to_nat. - -definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉. - -lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y. -#i #x #r #y % normalize - [cases (U i x r) normalize - [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H] - #H1 destruct - |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1] - #H1 // - ] - |#H >H //] -qed. - -lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?. -#i #x #r % normalize - [cases (U i x r) normalize // - #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1] - #H1 destruct - |#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)). - -lemma min_input_def : ∀h,i,x. - min_input h i x = min (x -i) (S i) (λy.termb i y (h (S i) y)). -// qed. - -lemma min_input_i: ∀h,i,x. x ≤ i → min_input h i x = S i. -#h #i #x #lexi >min_input_def -cut (x - i = 0) [@sym_eq /2 by eq_minus_O/] #Hcut // -qed. - -lemma min_input_to_terminate: ∀h,i,x. - min_input h i x = x → {i ⊙ x} ↓ (h (S i) x). -#h #i #x #Hminx -cases (decidable_le (S i) x) #Hix - [cases (true_or_false (termb i x (h (S i) x))) #Hcase - [@termb_true_to_term // - |min_input_def in Hminx; #Hminx >Hminx in ⊢ (%→?); - min_input_i in Hminx; - [#eqix >eqix in Hix; * /2/ | @le_S_S_to_le @not_le_to_lt //] - ] -qed. - -lemma min_input_to_lt: ∀h,i,x. - min_input h i x = x → i < x. -#h #i #x #Hminx cases (decidable_le (S i) x) // -#ltxi @False_ind >min_input_i in Hminx; - [#eqix >eqix in ltxi; * /2/ | @le_S_S_to_le @not_le_to_lt //] -qed. - -lemma le_to_min_input: ∀h,i,x,x1. x ≤ x1 → - min_input h i x = x → min_input h i x1 = x. -#h #i #x #x1 #lex #Hminx @(min_exists … (le_S_S … lex)) - [@(fmin_true … (sym_eq … Hminx)) // - |@(min_input_to_lt … Hminx) - |#j #H1 g_def cut (x-u = 0) [/2 by minus_le_minus_minus_comm/] -#eq0 >eq0 normalize // qed. - -lemma g_lt : ∀h,i,x. min_input h i x = x → - out i x (h (S i) x) < g h 0 x. -#h #i #x #H @le_S_S @(le_MaxI … i) /2 by min_input_to_lt/ -qed. - -lemma max_neq0 : ∀a,b. max a b ≠ 0 → a ≠ 0 ∨ b ≠ 0. -#a #b whd in match (max a b); cases (true_or_false (leb a b)) #Hcase >Hcase - [#H %2 @H | #H %1 @H] -qed. - -definition almost_equal ≝ λf,g:nat → nat. ¬ ∀nu.∃x. nu < x ∧ f x ≠ g x. -interpretation "almost equal" 'napart f g = (almost_equal f g). - -lemma eventually_cancelled: ∀h,u.¬∀nu.∃x. nu < x ∧ - max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) ≠ 0. -#h #u elim u - [normalize % #H cases (H u) #x * #_ * #H1 @H1 // - |#u0 @not_to_not #Hind #nu cases (Hind nu) #x * #ltx - cases (true_or_false (eqb (min_input h (u0+O) x) x)) #Hcase - [>bigop_Strue [2:@Hcase] #Hmax cases (max_neq0 … Hmax) -Hmax - [2: #H %{x} % // bigop_Sfalse - [#H %{x1} % [@transitive_lt //| (le_to_min_input … (eqb_true_to_eq … Hcase)) - [@lt_to_not_eq @ltx1 | @lt_to_le @ltx1] - ] - |>bigop_Sfalse [2:@Hcase] #H %{x} % // (bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA) - [>H // |@lt_to_le @(le_to_lt_to_lt …ltx) /2 by le_maxr/ |//] -qed. - -(******************************** Condition 2 *********************************) -definition total ≝ λf.λx:nat. Some nat (f x). - -lemma exists_to_exists_min: ∀h,i. (∃x. i < x ∧ {i ⊙ x} ↓ h (S i) x) → ∃y. min_input h i y = y. -#h #i * #x * #ltix #Hx %{(min_input h i x)} @min_spec_to_min @found // - [@(f_min_true (λy:ℕ.termb i y (h (S i) y))) %{x} % [% // | @term_to_termb_true //] - |#y #leiy #lty @(lt_min_to_false ????? lty) // - ] -qed. - -lemma condition_2: ∀h,i. code_for (total (g h 0)) i → ¬∃x. itermy >Hr -#H @(absurd ? H) @le_to_not_lt @le_n -qed. - - -(********************************* complexity *********************************) - -(* We assume operations have a minimal structural complexity MSC. -For instance, for time complexity, MSC is equal to the size of input. -For space complexity, MSC is typically 0, since we only measure the -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) *) - -definition C ≝ λs,i.∃c.∃a.∀x.a ≤ x → ∃y. - U i x (c*(s x)) = Some ? 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 ??); associative_times @le_times // @Ha1 @(transitive_le … lean) // - ] -qed. - -lemma timesc_CF: ∀s,f,c.CF (λx.c*s x) f → CF s f. -#s #f #c @O_to_CF @O_times_c -qed. - -(********************************* composition ********************************) -axiom CF_comp: ∀f,g,sf,sg,sh. CF sg g → CF sf f → - O sh (λx. sg x + sf (g x)) → CF sh (f ∘ g). - -lemma CF_comp_ext: ∀f,g,h,sh,sf,sg. CF sg g → CF sf f → - (∀x.f(g x) = h x) → O sh (λx. sg x + sf (g x)) → CF sh h. -#f #g #h #sh #sf #sg #Hg #Hf #Heq #H @(ext_CF (f ∘ g)) - [#n normalize @Heq | @(CF_comp … H) //] -qed. - - -(**************************** primitive operations*****************************) - -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_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 ***********************************) - -axiom CF_eqb: ∀h,f,g. - CF h f → CF h g → CF h (λx.eqb (f x) (g x)). - -(*********************************** maximum **********************************) - -axiom CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s. - CF ha a → CF hb b → CF hp p → CF hf f → - O s (λx.ha x + hb x + ∑_{i ∈[a x ,b x[ }(hp 〈i,x〉 + hf 〈i,x〉)) → - CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). - -(******************************** minimization ********************************) - -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〉)). - -(************************************* smn ************************************) -axiom smn: ∀f,s. CF s f → ∀x. CF (λy.s 〈x,y〉) (λy.f 〈x,y〉). - -(****************************** constructibility ******************************) - -definition constructible ≝ λs. CF s s. - -lemma constr_comp : ∀s1,s2. constructible s1 → constructible s2 → - (∀x. x ≤ s2 x) → constructible (s2 ∘ s1). -#s1 #s2 #Hs1 #Hs2 #Hle @(CF_comp … Hs1 Hs2) @O_plus @le_to_O #x [@Hle | //] -qed. - -lemma ext_constr: ∀s1,s2. (∀x.s1 x = s2 x) → - constructible s1 → constructible s2. -#s1 #s2 #Hext #Hs1 @(ext_CF … Hext) @(monotonic_CF … Hs1) #x >Hext // -qed. - -(********************************* simulation *********************************) - -axiom sU : nat → nat. - -axiom monotonic_sU: ∀i1,i2,x1,x2,s1,s2. i1 ≤ i2 → x1 ≤ x2 → s1 ≤ s2 → - sU 〈i1,〈x1,s1〉〉 ≤ sU 〈i2,〈x2,s2〉〉. - -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) -#b1 * #c1 #eqy >eqy -eqy -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〉〉. - -definition pU_unary ≝ λp. pU (fst p) (fst (snd p)) (snd (snd p)). - -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. - - -(******************** complexity of g ********************) - -definition unary_g ≝ λh.λux. g h (fst ux) (snd ux). -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 (auxg h) → CF s (unary_g h). -#h #s #H1 @(CF_compS ? (auxg h) H1) -qed. - -definition aux1g ≝ - λh,ux. max_{i ∈[fst ux,snd ux[ | (λp. eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) 〈i,ux〉} - ((λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) 〈i,ux〉). - -lemma eq_aux : ∀h,x.aux1g h x = auxg h x. -#h #x @same_bigop - [#n #_ >fst_pair >snd_pair // |#n #_ #_ >fst_pair >snd_pair //] -qed. - -lemma compl_g2 : ∀h,s1,s2,s. - CF s1 - (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) → - CF s2 - (λ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 (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 (λ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 min_input_def whd in ⊢ (??%?); >minus_S_S @min_f_g #i #_ #_ -whd in ⊢ (??%%); >fst_pair >snd_pair // -qed. - -definition termb_aux ≝ λh. - termb_unary ∘ λp.〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉. - -lemma compl_g4 : ∀h,s1,s. - (CF s1 - (λ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 (λ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 - |@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //] -qed. - -(************************* a couple of technical lemmas ***********************) -lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0. -#a elim a // #n #Hind * - [#H @False_ind /2 by absurd/ | #b normalize #H @Hind @le_S_S_to_le /2/] -qed. - -lemma sigma_bound: ∀h,a,b. monotonic nat le h → - ∑_{i ∈ [a,S b[ }(h i) ≤ (S b-a)*h b. -#h #a #b #H cases (decidable_le a b) - [#leab cut (b = pred (S b - a + a)) - [Hb in match (h b); - generalize in match (S b -a); - #n elim n - [// - |#m #Hind >bigop_Strue [2://] @le_plus - [@H @le_n |@(transitive_le … Hind) @le_times [//] @H //] - ] - |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba - cut (S b -a = 0) [@minus_to_0 //] #Hcut >Hcut // - ] -qed. - -lemma sigma_bound_decr: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) → - ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a. -#h #a #b #H cases (decidable_le a b) - [#leab cut ((b -a) +a ≤ b) [/2 by le_minus_to_plus_r/] generalize in match (b -a); - #n elim n - [// - |#m #Hind >bigop_Strue [2://] #Hm - cut (m+a ≤ b) [@(transitive_le … Hm) //] #Hm1 - @le_plus [@H // |@(transitive_le … (Hind Hm1)) //] - ] - |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba - cut (b -a = 0) [@minus_to_0 @lt_to_le @ltba] #Hcut >Hcut // - ] -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〉)). -#s1 #Hs1 %{1} %{0} #n #_ >commutative_times minus_S_S //] -qed. - -lemma coroll2: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) → -O (λx.(snd x - fst x)*s1 〈fst x,x〉) (λx.∑_{i ∈[fst x,snd x[ }(s1 〈i,x〉)). -#s1 #Hs1 %{1} %{0} #n #_ >commutative_times fst_pair >snd_pair >fst_pair >snd_pair // ] -@(CF_comp … (λx.MSC 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 - [@O_plus - [@(O_trans … (λx.MSC (fst x) + MSC (max (fst (snd x)) (snd (snd x))))) - [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx - >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) - >distributive_times_plus @le_plus [//] - cases (surj_pair b) #c * #d #eqb >eqb - >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) - whd in ⊢ (??%); @le_plus - [@monotonic_MSC @(le_maxl … (le_n …)) - |>commutative_times fst_pair >snd_pair // qed. - -lemma le_big : ∀x. x ≤ big x. -#x cases (surj_pair x) #a * #b #eqx >eqx @le_pair >fst_pair >snd_pair -[@(le_maxl … (le_n …)) | @(le_maxr … (le_n …))] -qed. - -definition faux2 ≝ λh. - (λx.MSC x + (snd (snd x)-fst x)* - (λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉). - -lemma compl_g7: ∀h. - constructible (λx. h (fst x) (snd x)) → - (∀n. monotonic ? le (h n)) → - CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉) - (λp:ℕ.min_input h (fst p) (snd (snd p))). -#h #hcostr #hmono @(monotonic_CF … (faux2 h)) - [#n normalize >fst_pair >snd_pair //] -@compl_g5 [2:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair ->fst_pair >snd_pair @monotonic_sU // @hmono @lexy -qed. - -lemma compl_g71: ∀h. - 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))〉〉) - (λp:ℕ.min_input h (fst p) (snd (snd p))). -#h #hcostr #hmono @(monotonic_CF … (compl_g7 h hcostr hmono)) #x -@le_plus [@monotonic_MSC //] -cases (decidable_le (fst x) (snd(snd x))) - [#Hle @le_times // @monotonic_sU - |#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt] - ] -qed. - -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 (MSC x)) >commutative_times @le_times - [// | @monotonic_MSC // ]] -@(O_trans … (coroll2 ??)) - [#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 [//|@(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 - [/2 by monotonic_le_minus_r/ - |@monotonic_sU // @hantimono [@le_S_S // |@ltb] - ] - |@monotonic_sU // @hantimono [@le_S_S // |@ltb] - ] - |@le_to_O #n >fst_pair >snd_pair - cut (max (fst n) n = n) [normalize >le_to_leb_true //] #Hmax >Hmax - >associative_plus >distributive_times_plus - @le_plus [@le_times [@le_S // |>big_def >Hmax //] |//] - ] -qed. - -definition sg ≝ λh,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)〉〉. - -lemma sg_def : ∀h,a,b. - sg h 〈a,b〉 = (S (b-a))*MSC 〈〈a,b〉,〈a,b〉〉 + - (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉. -#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair // -qed. - -lemma compl_g11 : ∀h. - 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) (unary_g h). -#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham) -qed. - -(**************************** closing the argument ****************************) - -let rec h_of_aux (r:nat →nat) (c,d,b:nat) on d : nat ≝ - match d with - [ O ⇒ c - | S d1 ⇒ (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + - d*(S d)*sU 〈〈b-d,b〉,〈b,r (h_of_aux r c d1 b)〉〉]. - -lemma h_of_aux_O: ∀r,c,b. - h_of_aux r c O b = c. -// qed. - -lemma h_of_aux_S : ∀r,c,d,b. - h_of_aux r c (S d) b = - (S (S d))*(MSC 〈〈b-(S d),b〉,〈b-(S d),b〉〉) + - (S d)*(S (S d))*sU 〈〈b-(S d),b〉,〈b,r(h_of_aux r c d b)〉〉. -// qed. - -definition h_of ≝ λr,p. - let m ≝ max (fst p) (snd p) in - h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (snd p - fst p) (snd p). - -lemma h_of_O: ∀r,a,b. b ≤ a → - h_of r 〈a,b〉 = let m ≝ max a b in MSC 〈〈m,m〉,〈m,m〉〉. -#r #a #b #Hle normalize >fst_pair >snd_pair >(minus_to_0 … Hle) // -qed. - -lemma h_of_def: ∀r,a,b.h_of r 〈a,b〉 = - let m ≝ max a b in - h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (b - a) b. -#r #a #b normalize >fst_pair >snd_pair // -qed. - -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. -#r #Hr #monor #d #d1 lapply d -d elim d1 - [#d #c #c1 #b #b1 #Hc #Hd @(le_n_O_elim ? Hd) #leb - >h_of_aux_O >h_of_aux_O // - |#m #Hind #d #c #c1 #b #b1 #lec #led #leb cases (le_to_or_lt_eq … led) - [#ltd @(transitive_le … (Hind … lec ? leb)) [@le_S_S_to_le @ltd] - >h_of_aux_S @(transitive_le ???? (le_plus_n …)) - >(times_n_1 (h_of_aux r c1 m b1)) in ⊢ (?%?); - >commutative_times @le_times [//|@(transitive_le … (Hr ?)) @sU_le] - |#Hd >Hd >h_of_aux_S >h_of_aux_S - cut (b-S m ≤ b1 - S m) [/2 by monotonic_le_minus_l/] #Hb1 - @le_plus [@le_times //] - [@monotonic_MSC @le_pair @le_pair // - |@le_times [//] @monotonic_sU - [@le_pair // |// |@monor @Hind //] - ] - ] - ] -qed. - -lemma mono_h_of2: ∀r.(∀x. x ≤ r x) → monotonic ? le r → - ∀i,b,b1. b ≤ b1 → h_of r 〈i,b〉 ≤ h_of r 〈i,b1〉. -#r #Hr #Hmono #i #a #b #leab >h_of_def >h_of_def -cut (max i a ≤ max i b) - [@to_max - [@(le_maxl … (le_n …))|@(transitive_le … leab) @(le_maxr … (le_n …))]] -#Hmax @(mono_h_of_aux r Hr Hmono) - [@monotonic_MSC @le_pair @le_pair @Hmax |/2 by monotonic_le_minus_l/ |@leab] -qed. - -axiom h_of_constr : ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - constructible (h_of r). - -lemma speed_compl: ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))). -#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …)) - [#x cases (surj_pair x) #a * #b #eqx >eqx - >sg_def cases (decidable_le b a) - [#leba >(minus_to_0 … leba) normalize in ⊢ (?%?); - h_of_def - cut (max a b = a) - [normalize cases (le_to_or_lt_eq … leba) - [#ltba >(lt_to_leb_false … ltba) % - |#eqba (le_to_leb_true … (le_n ?)) % ]] - #Hmax >Hmax normalize >(minus_to_0 … leba) normalize - @monotonic_MSC @le_pair @le_pair // - |#ltab >h_of_def >h_of_def - cut (max a b = b) - [normalize >(le_to_leb_true … ) [%] @lt_to_le @not_le_to_lt @ltab] - #Hmax >Hmax - 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) - [%{(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 - [@lt_to_le @not_le_to_lt // | //]] #eqd2 >eqd2 - normalize // - ] - |#n #a #b #leab #lebn >h_of_def >h_of_def - cut (max a n = n) - [normalize >le_to_leb_true [%|@(transitive_le … leab lebn)]] #Hmaxa - cut (max b n = n) - [normalize >(le_to_leb_true … lebn) %] #Hmaxb - >Hmaxa >Hmaxb @Hmono @(mono_h_of_aux r … Hr Hmono) // /2 by monotonic_le_minus_r/ - |#n #a #b #leab @Hmono @(mono_h_of2 … Hr Hmono … leab) - |@(constr_comp … Hconstr Hr) @(ext_constr (h_of r)) - [#x cases (surj_pair x) #a * #b #eqx >eqx >fst_pair >snd_pair //] - @(h_of_constr r Hr Hmono Hconstr) - ] -qed. - -lemma speed_compl_i: ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∀i. CF (λx.h_of r 〈i,x〉) (λx.g (λi,x. r(h_of r 〈i,x〉)) i x). -#r #Hr #Hmono #Hconstr #i -@(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 Hconstr)) #n // -qed. - -(**************************** the speedup theorem *****************************) -theorem pseudo_speedup: - ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∃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 #Hconstr -(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * -#Hcodei #HCi -(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} -(* sg is (λx.h_of r 〈i,x〉) *) -%{(λx. h_of r 〈S i,x〉)} -lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg -%[%[@condition_1 |@Hg] - |cases Hg #H1 * #j * #Hcodej #HCj - lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) - cases HCi #m * #a #Ha %{m} %{(max (S i) a)} #n #ltin @lt_to_le @not_le_to_lt - @(not_to_not … Hcond2) -Hcond2 #Hlesf %{n} % - [@(transitive_le … ltin) @(le_maxl … (le_n …))] - cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] - #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) // - ] -qed. - -theorem pseudo_speedup': - ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∃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 #Hconstr -(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * -#Hcodei #HCi -(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} -(* sg is (λx.h_of r 〈i,x〉) *) -%{(λx. h_of r 〈S i,x〉)} -lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg -%[%[@condition_1 |@Hg] - |cases Hg #H1 * #j * #Hcodej #HCj - lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) - cases HCi #m * #a #Ha - %{m} %{(max (S i) a)} #n #ltin @not_le_to_lt @(not_to_not … Hcond2) -Hcond2 #Hlesf - %{n} % [@(transitive_le … ltin) @(le_maxl … (le_n …))] - cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] - #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) - @Hmono @(mono_h_of2 … Hr Hmono … ltin) - ] -qed. - diff --git a/matita/matita/broken_lib/reverse_complexity/speed_new.ma b/matita/matita/broken_lib/reverse_complexity/speed_new.ma deleted file mode 100644 index 9812cfb08..000000000 --- a/matita/matita/broken_lib/reverse_complexity/speed_new.ma +++ /dev/null @@ -1,922 +0,0 @@ -include "basics/types.ma". -include "arithmetics/minimization.ma". -include "arithmetics/bigops.ma". -include "arithmetics/sigma_pi.ma". -include "arithmetics/bounded_quantifiers.ma". -include "reverse_complexity/big_O.ma". -include "basics/core_notation/napart_2.ma". - -(************************* notation for minimization *****************************) -notation "μ_{ ident i < n } p" - with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. - -notation "μ_{ ident i ≤ n } p" - with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. - -notation "μ_{ ident i ∈ [a,b[ } p" - with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}. - -notation "μ_{ ident i ∈ [a,b] } p" - with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. - -(************************************ 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 le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true → - f a ≤ Max_{i ∈ [m,n[ | p i}(f i). -#f #p #n #m #a #lema #ltan #pa ->(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m)) - [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. - -(********************************** pairing ***********************************) -axiom pair: nat → nat → nat. -axiom fst : nat → nat. -axiom snd : nat → nat. - -interpretation "abstract pair" 'pair f g = (pair f g). - -axiom fst_pair: ∀a,b. fst 〈a,b〉 = a. -axiom snd_pair: ∀a,b. snd 〈a,b〉 = b. -axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉. - -axiom le_fst : ∀p. fst p ≤ p. -axiom le_snd : ∀p. snd p ≤ p. -axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉. - -(************************************* U **************************************) -axiom U: nat → nat →nat → option nat. - -axiom monotonic_U: ∀i,x,n,m,y.n ≤m → - U i x n = Some ? y → U i x m = Some ? y. - -lemma unique_U: ∀i,x,n,m,yn,ym. - U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. -#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. - -notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. - -lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n. -#i #x #n normalize cases (U i x n) - [%2 % * #y #H destruct|#y %1 %{y} //] -qed. - -lemma monotonic_terminate: ∀i,x,n,m. - n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m. -#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) // -qed. - -definition termb ≝ λi,x,t. - match U i x t with [None ⇒ false |Some y ⇒ true]. - -lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t. -#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] -qed. - -lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true. -#i #x #t * #y #H normalize >H // -qed. - -definition out ≝ λi,x,r. - match U i x r with [ None ⇒ 0 | Some z ⇒ z]. - -definition bool_to_nat: bool → nat ≝ - λb. match b with [true ⇒ 1 | false ⇒ 0]. - -coercion bool_to_nat. - -definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉. - -lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y. -#i #x #r #y % normalize - [cases (U i x r) normalize - [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H] - #H1 destruct - |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1] - #H1 // - ] - |#H >H //] -qed. - -lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?. -#i #x #r % normalize - [cases (U i x r) normalize // - #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1] - #H1 destruct - |#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)). - -lemma min_input_def : ∀h,i,x. - min_input h i x = min (x -i) (S i) (λy.termb i y (h (S i) y)). -// qed. - -lemma min_input_i: ∀h,i,x. x ≤ i → min_input h i x = S i. -#h #i #x #lexi >min_input_def -cut (x - i = 0) [@sym_eq /2 by eq_minus_O/] #Hcut // -qed. - -lemma min_input_to_terminate: ∀h,i,x. - min_input h i x = x → {i ⊙ x} ↓ (h (S i) x). -#h #i #x #Hminx -cases (decidable_le (S i) x) #Hix - [cases (true_or_false (termb i x (h (S i) x))) #Hcase - [@termb_true_to_term // - |min_input_def in Hminx; #Hminx >Hminx in ⊢ (%→?); - min_input_i in Hminx; - [#eqix >eqix in Hix; * /2/ | @le_S_S_to_le @not_le_to_lt //] - ] -qed. - -lemma min_input_to_lt: ∀h,i,x. - min_input h i x = x → i < x. -#h #i #x #Hminx cases (decidable_le (S i) x) // -#ltxi @False_ind >min_input_i in Hminx; - [#eqix >eqix in ltxi; * /2/ | @le_S_S_to_le @not_le_to_lt //] -qed. - -lemma le_to_min_input: ∀h,i,x,x1. x ≤ x1 → - min_input h i x = x → min_input h i x1 = x. -#h #i #x #x1 #lex #Hminx @(min_exists … (le_S_S … lex)) - [@(fmin_true … (sym_eq … Hminx)) // - |@(min_input_to_lt … Hminx) - |#j #H1 g_def cut (x-u = 0) [/2 by minus_le_minus_minus_comm/] -#eq0 >eq0 normalize // qed. - -lemma g_lt : ∀h,i,x. min_input h i x = x → - out i x (h (S i) x) < g h 0 x. -#h #i #x #H @le_S_S @(le_MaxI … i) /2 by min_input_to_lt/ -qed. - -lemma max_neq0 : ∀a,b. max a b ≠ 0 → a ≠ 0 ∨ b ≠ 0. -#a #b whd in match (max a b); cases (true_or_false (leb a b)) #Hcase >Hcase - [#H %2 @H | #H %1 @H] -qed. - -definition almost_equal ≝ λf,g:nat → nat. ¬ ∀nu.∃x. nu < x ∧ f x ≠ g x. -interpretation "almost equal" 'napart f g = (almost_equal f g). - -lemma eventually_cancelled: ∀h,u.¬∀nu.∃x. nu < x ∧ - max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) ≠ 0. -#h #u elim u - [normalize % #H cases (H u) #x * #_ * #H1 @H1 // - |#u0 @not_to_not #Hind #nu cases (Hind nu) #x * #ltx - cases (true_or_false (eqb (min_input h (u0+O) x) x)) #Hcase - [>bigop_Strue [2:@Hcase] #Hmax cases (max_neq0 … Hmax) -Hmax - [2: #H %{x} % // bigop_Sfalse - [#H %{x1} % [@transitive_lt //| (le_to_min_input … (eqb_true_to_eq … Hcase)) - [@lt_to_not_eq @ltx1 | @lt_to_le @ltx1] - ] - |>bigop_Sfalse [2:@Hcase] #H %{x} % // (bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA) - [>H // |@lt_to_le @(le_to_lt_to_lt …ltx) /2 by le_maxr/ |//] -qed. - -(******************************** Condition 2 *********************************) -definition total ≝ λf.λx:nat. Some nat (f x). - -lemma exists_to_exists_min: ∀h,i. (∃x. i < x ∧ {i ⊙ x} ↓ h (S i) x) → ∃y. min_input h i y = y. -#h #i * #x * #ltix #Hx %{(min_input h i x)} @min_spec_to_min @found // - [@(f_min_true (λy:ℕ.termb i y (h (S i) y))) %{x} % [% // | @term_to_termb_true //] - |#y #leiy #lty @(lt_min_to_false ????? lty) // - ] -qed. - -lemma condition_2: ∀h,i. code_for (total (g h 0)) i → ¬∃x. itermy >Hr -#H @(absurd ? H) @le_to_not_lt @le_n -qed. - - -(********************************* complexity *********************************) - -(* We assume operations have a minimal structural complexity MSC. -For instance, for time complexity, MSC is equal to the size of input. -For space complexity, MSC is typically 0, since we only measure the -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) *) - -definition C ≝ λs,i.∃c.∃a.∀x.a ≤ x → ∃y. - U i x (c*(s x)) = Some ? 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 ??); associative_times @le_times // @Ha1 @(transitive_le … lean) // - ] -qed. - -lemma timesc_CF: ∀s,f,c.CF (λx.c*s x) f → CF s f. -#s #f #c @O_to_CF @O_times_c -qed. - -(********************************* composition ********************************) -axiom CF_comp: ∀f,g,sf,sg,sh. CF sg g → CF sf f → - O sh (λx. sg x + sf (g x)) → CF sh (f ∘ g). - -lemma CF_comp_ext: ∀f,g,h,sh,sf,sg. CF sg g → CF sf f → - (∀x.f(g x) = h x) → O sh (λx. sg x + sf (g x)) → CF sh h. -#f #g #h #sh #sf #sg #Hg #Hf #Heq #H @(ext_CF (f ∘ g)) - [#n normalize @Heq | @(CF_comp … H) //] -qed. - - -(**************************** primitive operations*****************************) - -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_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 ***********************************) - -axiom CF_eqb: ∀h,f,g. - CF h f → CF h g → CF h (λx.eqb (f x) (g x)). - -(*********************************** maximum **********************************) - -axiom CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s. - CF ha a → CF hb b → CF hp p → CF hf f → - O s (λx.ha x + hb x + ∑_{i ∈[a x ,b x[ }(hp 〈i,x〉 + hf 〈i,x〉)) → - CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). - -(******************************** minimization ********************************) - -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〉)). - -(************************************* smn ************************************) -axiom smn: ∀f,s. CF s f → ∀x. CF (λy.s 〈x,y〉) (λy.f 〈x,y〉). - -(****************************** constructibility ******************************) - -definition constructible ≝ λs. CF s s. - -lemma constr_comp : ∀s1,s2. constructible s1 → constructible s2 → - (∀x. x ≤ s2 x) → constructible (s2 ∘ s1). -#s1 #s2 #Hs1 #Hs2 #Hle @(CF_comp … Hs1 Hs2) @O_plus @le_to_O #x [@Hle | //] -qed. - -lemma ext_constr: ∀s1,s2. (∀x.s1 x = s2 x) → - constructible s1 → constructible s2. -#s1 #s2 #Hext #Hs1 @(ext_CF … Hext) @(monotonic_CF … Hs1) #x >Hext // -qed. - -(********************************* simulation *********************************) - -axiom sU : nat → nat. - -axiom monotonic_sU: ∀i1,i2,x1,x2,s1,s2. i1 ≤ i2 → x1 ≤ x2 → s1 ≤ s2 → - sU 〈i1,〈x1,s1〉〉 ≤ sU 〈i2,〈x2,s2〉〉. - -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) -#b1 * #c1 #eqy >eqy -eqy -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〉〉. - -definition pU_unary ≝ λp. pU (fst p) (fst (snd p)) (snd (snd p)). - -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. - - -(******************** complexity of g ********************) - -definition unary_g ≝ λh.λux. g h (fst ux) (snd ux). -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 (auxg h) → CF s (unary_g h). -#h #s #H1 @(CF_compS ? (auxg h) H1) -qed. - -definition aux1g ≝ - λh,ux. max_{i ∈[fst ux,snd ux[ | (λp. eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) 〈i,ux〉} - ((λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) 〈i,ux〉). - -lemma eq_aux : ∀h,x.aux1g h x = auxg h x. -#h #x @same_bigop - [#n #_ >fst_pair >snd_pair // |#n #_ #_ >fst_pair >snd_pair //] -qed. - -lemma compl_g2 : ∀h,s1,s2,s. - CF s1 - (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) → - CF s2 - (λ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 (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 (λ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 min_input_def whd in ⊢ (??%?); >minus_S_S @min_f_g #i #_ #_ -whd in ⊢ (??%%); >fst_pair >snd_pair // -qed. - -definition termb_aux ≝ λh. - termb_unary ∘ λp.〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉. - -lemma compl_g4 : ∀h,s1,s. - (CF s1 - (λ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 (λ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 - |@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //] -qed. - -(************************* a couple of technical lemmas ***********************) -lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0. -#a elim a // #n #Hind * - [#H @False_ind /2 by absurd/ | #b normalize #H @Hind @le_S_S_to_le /2/] -qed. - -lemma sigma_bound: ∀h,a,b. monotonic nat le h → - ∑_{i ∈ [a,S b[ }(h i) ≤ (S b-a)*h b. -#h #a #b #H cases (decidable_le a b) - [#leab cut (b = pred (S b - a + a)) - [Hb in match (h b); - generalize in match (S b -a); - #n elim n - [// - |#m #Hind >bigop_Strue [2://] @le_plus - [@H @le_n |@(transitive_le … Hind) @le_times [//] @H //] - ] - |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba - cut (S b -a = 0) [@minus_to_0 //] #Hcut >Hcut // - ] -qed. - -lemma sigma_bound_decr: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) → - ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a. -#h #a #b #H cases (decidable_le a b) - [#leab cut ((b -a) +a ≤ b) [/2 by le_minus_to_plus_r/] generalize in match (b -a); - #n elim n - [// - |#m #Hind >bigop_Strue [2://] #Hm - cut (m+a ≤ b) [@(transitive_le … Hm) //] #Hm1 - @le_plus [@H // |@(transitive_le … (Hind Hm1)) //] - ] - |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba - cut (b -a = 0) [@minus_to_0 @lt_to_le @ltba] #Hcut >Hcut // - ] -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〉)). -#s1 #Hs1 %{1} %{0} #n #_ >commutative_times minus_S_S //] -qed. - -lemma coroll2: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) → -O (λx.(snd x - fst x)*s1 〈fst x,x〉) (λx.∑_{i ∈[fst x,snd x[ }(s1 〈i,x〉)). -#s1 #Hs1 %{1} %{0} #n #_ >commutative_times fst_pair >snd_pair >fst_pair >snd_pair // ] -@(CF_comp … (λx.MSC 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 - [@O_plus - [@(O_trans … (λx.MSC (fst x) + MSC (max (fst (snd x)) (snd (snd x))))) - [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx - >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) - >distributive_times_plus @le_plus [//] - cases (surj_pair b) #c * #d #eqb >eqb - >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) - whd in ⊢ (??%); @le_plus - [@monotonic_MSC @(le_maxl … (le_n …)) - |>commutative_times fst_pair >snd_pair // qed. - -lemma le_big : ∀x. x ≤ big x. -#x cases (surj_pair x) #a * #b #eqx >eqx @le_pair >fst_pair >snd_pair -[@(le_maxl … (le_n …)) | @(le_maxr … (le_n …))] -qed. - -definition faux2 ≝ λh. - (λx.MSC x + (snd (snd x)-fst x)* - (λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉). - -lemma compl_g7: ∀h. - constructible (λx. h (fst x) (snd x)) → - (∀n. monotonic ? le (h n)) → - CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉) - (λp:ℕ.min_input h (fst p) (snd (snd p))). -#h #hcostr #hmono @(monotonic_CF … (faux2 h)) - [#n normalize >fst_pair >snd_pair //] -@compl_g5 [2:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair ->fst_pair >snd_pair @monotonic_sU // @hmono @lexy -qed. - -lemma compl_g71: ∀h. - 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))〉〉) - (λp:ℕ.min_input h (fst p) (snd (snd p))). -#h #hcostr #hmono @(monotonic_CF … (compl_g7 h hcostr hmono)) #x -@le_plus [@monotonic_MSC //] -cases (decidable_le (fst x) (snd(snd x))) - [#Hle @le_times // @monotonic_sU - |#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt] - ] -qed. - -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 (MSC x)) >commutative_times @le_times - [// | @monotonic_MSC // ]] -@(O_trans … (coroll2 ??)) - [#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 [//|@(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 - [/2 by monotonic_le_minus_r/ - |@monotonic_sU // @hantimono [@le_S_S // |@ltb] - ] - |@monotonic_sU // @hantimono [@le_S_S // |@ltb] - ] - |@le_to_O #n >fst_pair >snd_pair - cut (max (fst n) n = n) [normalize >le_to_leb_true //] #Hmax >Hmax - >associative_plus >distributive_times_plus - @le_plus [@le_times [@le_S // |>big_def >Hmax //] |//] - ] -qed. - -definition sg ≝ λh,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)〉〉. - -lemma sg_def : ∀h,a,b. - sg h 〈a,b〉 = (S (b-a))*MSC 〈〈a,b〉,〈a,b〉〉 + - (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉. -#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair // -qed. - -lemma compl_g11 : ∀h. - 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) (unary_g h). -#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham) -qed. - -(**************************** closing the argument ****************************) - -let rec h_of_aux (r:nat →nat) (c,d,b:nat) on d : nat ≝ - match d with - [ O ⇒ c - | S d1 ⇒ (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + - d*(S d)*sU 〈〈b-d,b〉,〈b,r (h_of_aux r c d1 b)〉〉]. - -lemma h_of_aux_O: ∀r,c,b. - h_of_aux r c O b = c. -// qed. - -lemma h_of_aux_S : ∀r,c,d,b. - h_of_aux r c (S d) b = - (S (S d))*(MSC 〈〈b-(S d),b〉,〈b-(S d),b〉〉) + - (S d)*(S (S d))*sU 〈〈b-(S d),b〉,〈b,r(h_of_aux r c d b)〉〉. -// qed. - -definition h_of ≝ λr,p. - let m ≝ max (fst p) (snd p) in - h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (snd p - fst p) (snd p). - -lemma h_of_O: ∀r,a,b. b ≤ a → - h_of r 〈a,b〉 = let m ≝ max a b in MSC 〈〈m,m〉,〈m,m〉〉. -#r #a #b #Hle normalize >fst_pair >snd_pair >(minus_to_0 … Hle) // -qed. - -lemma h_of_def: ∀r,a,b.h_of r 〈a,b〉 = - let m ≝ max a b in - h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (b - a) b. -#r #a #b normalize >fst_pair >snd_pair // -qed. - -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. -#r #Hr #monor #d #d1 lapply d -d elim d1 - [#d #c #c1 #b #b1 #Hc #Hd @(le_n_O_elim ? Hd) #leb - >h_of_aux_O >h_of_aux_O // - |#m #Hind #d #c #c1 #b #b1 #lec #led #leb cases (le_to_or_lt_eq … led) - [#ltd @(transitive_le … (Hind … lec ? leb)) [@le_S_S_to_le @ltd] - >h_of_aux_S @(transitive_le ???? (le_plus_n …)) - >(times_n_1 (h_of_aux r c1 m b1)) in ⊢ (?%?); - >commutative_times @le_times [//|@(transitive_le … (Hr ?)) @sU_le] - |#Hd >Hd >h_of_aux_S >h_of_aux_S - cut (b-S m ≤ b1 - S m) [/2 by monotonic_le_minus_l/] #Hb1 - @le_plus [@le_times //] - [@monotonic_MSC @le_pair @le_pair // - |@le_times [//] @monotonic_sU - [@le_pair // |// |@monor @Hind //] - ] - ] - ] -qed. - -lemma mono_h_of2: ∀r.(∀x. x ≤ r x) → monotonic ? le r → - ∀i,b,b1. b ≤ b1 → h_of r 〈i,b〉 ≤ h_of r 〈i,b1〉. -#r #Hr #Hmono #i #a #b #leab >h_of_def >h_of_def -cut (max i a ≤ max i b) - [@to_max - [@(le_maxl … (le_n …))|@(transitive_le … leab) @(le_maxr … (le_n …))]] -#Hmax @(mono_h_of_aux r Hr Hmono) - [@monotonic_MSC @le_pair @le_pair @Hmax |/2 by monotonic_le_minus_l/ |@leab] -qed. - -axiom h_of_constr : ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - constructible (h_of r). - -lemma speed_compl: ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))). -#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …)) - [#x cases (surj_pair x) #a * #b #eqx >eqx - >sg_def cases (decidable_le b a) - [#leba >(minus_to_0 … leba) normalize in ⊢ (?%?); - h_of_def - cut (max a b = a) - [normalize cases (le_to_or_lt_eq … leba) - [#ltba >(lt_to_leb_false … ltba) % - |#eqba (le_to_leb_true … (le_n ?)) % ]] - #Hmax >Hmax normalize >(minus_to_0 … leba) normalize - @monotonic_MSC @le_pair @le_pair // - |#ltab >h_of_def >h_of_def - cut (max a b = b) - [normalize >(le_to_leb_true … ) [%] @lt_to_le @not_le_to_lt @ltab] - #Hmax >Hmax - 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) - [%{(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 - [@lt_to_le @not_le_to_lt // | //]] #eqd2 >eqd2 - normalize // - ] - |#n #a #b #leab #lebn >h_of_def >h_of_def - cut (max a n = n) - [normalize >le_to_leb_true [%|@(transitive_le … leab lebn)]] #Hmaxa - cut (max b n = n) - [normalize >(le_to_leb_true … lebn) %] #Hmaxb - >Hmaxa >Hmaxb @Hmono @(mono_h_of_aux r … Hr Hmono) // /2 by monotonic_le_minus_r/ - |#n #a #b #leab @Hmono @(mono_h_of2 … Hr Hmono … leab) - |@(constr_comp … Hconstr Hr) @(ext_constr (h_of r)) - [#x cases (surj_pair x) #a * #b #eqx >eqx >fst_pair >snd_pair //] - @(h_of_constr r Hr Hmono Hconstr) - ] -qed. - -lemma speed_compl_i: ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∀i. CF (λx.h_of r 〈i,x〉) (λx.g (λi,x. r(h_of r 〈i,x〉)) i x). -#r #Hr #Hmono #Hconstr #i -@(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 Hconstr)) #n // -qed. - -(**************************** the speedup theorem *****************************) -theorem pseudo_speedup: - ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∃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 #Hconstr -(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * -#Hcodei #HCi -(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} -(* sg is (λx.h_of r 〈i,x〉) *) -%{(λx. h_of r 〈S i,x〉)} -lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg -%[%[@condition_1 |@Hg] - |cases Hg #H1 * #j * #Hcodej #HCj - lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) - cases HCi #m * #a #Ha %{m} %{(max (S i) a)} #n #ltin @lt_to_le @not_le_to_lt - @(not_to_not … Hcond2) -Hcond2 #Hlesf %{n} % - [@(transitive_le … ltin) @(le_maxl … (le_n …))] - cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] - #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) // - ] -qed. - -theorem pseudo_speedup': - ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∃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 #Hconstr -(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * -#Hcodei #HCi -(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} -(* sg is (λx.h_of r 〈i,x〉) *) -%{(λx. h_of r 〈S i,x〉)} -lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg -%[%[@condition_1 |@Hg] - |cases Hg #H1 * #j * #Hcodej #HCj - lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) - cases HCi #m * #a #Ha - %{m} %{(max (S i) a)} #n #ltin @not_le_to_lt @(not_to_not … Hcond2) -Hcond2 #Hlesf - %{n} % [@(transitive_le … ltin) @(le_maxl … (le_n …))] - cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] - #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) - @Hmono @(mono_h_of2 … Hr Hmono … ltin) - ] -qed. - diff --git a/matita/matita/broken_lib/reverse_complexity/toolkit.ma b/matita/matita/broken_lib/reverse_complexity/toolkit.ma deleted file mode 100644 index 69fa3525a..000000000 --- a/matita/matita/broken_lib/reverse_complexity/toolkit.ma +++ /dev/null @@ -1,988 +0,0 @@ -include "basics/types.ma". -include "arithmetics/minimization.ma". -include "arithmetics/bigops.ma". -include "arithmetics/sigma_pi.ma". -include "arithmetics/bounded_quantifiers.ma". -include "reverse_complexity/big_O.ma". -include "basics/core_notation/napart_2.ma". - -(************************* notation for minimization *****************************) -notation "μ_{ ident i < n } p" - with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. - -notation "μ_{ ident i ≤ n } p" - with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. - -notation "μ_{ ident i ∈ [a,b[ } p" - with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}. - -notation "μ_{ ident i ∈ [a,b] } p" - with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. - -(************************************ 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 le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true → - f a ≤ Max_{i ∈ [m,n[ | p i}(f i). -#f #p #n #m #a #lema #ltan #pa ->(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m)) - [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. - -(********************************** pairing ***********************************) -axiom pair: nat → nat → nat. -axiom fst : nat → nat. -axiom snd : nat → nat. - -interpretation "abstract pair" 'pair f g = (pair f g). - -axiom fst_pair: ∀a,b. fst 〈a,b〉 = a. -axiom snd_pair: ∀a,b. snd 〈a,b〉 = b. -axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉. - -axiom le_fst : ∀p. fst p ≤ p. -axiom le_snd : ∀p. snd p ≤ p. -axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉. - -(************************************* U **************************************) -axiom U: nat → nat →nat → option nat. - -axiom monotonic_U: ∀i,x,n,m,y.n ≤m → - U i x n = Some ? y → U i x m = Some ? y. - -lemma unique_U: ∀i,x,n,m,yn,ym. - U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. -#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. - -notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. - -lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n. -#i #x #n normalize cases (U i x n) - [%2 % * #y #H destruct|#y %1 %{y} //] -qed. - -lemma monotonic_terminate: ∀i,x,n,m. - n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m. -#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) // -qed. - -definition termb ≝ λi,x,t. - match U i x t with [None ⇒ false |Some y ⇒ true]. - -lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t. -#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] -qed. - -lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true. -#i #x #t * #y #H normalize >H // -qed. - -definition out ≝ λi,x,r. - match U i x r with [ None ⇒ 0 | Some z ⇒ z]. - -definition bool_to_nat: bool → nat ≝ - λb. match b with [true ⇒ 1 | false ⇒ 0]. - -coercion bool_to_nat. - -definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉. - -lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y. -#i #x #r #y % normalize - [cases (U i x r) normalize - [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H] - #H1 destruct - |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1] - #H1 // - ] - |#H >H //] -qed. - -lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?. -#i #x #r % normalize - [cases (U i x r) normalize // - #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1] - #H1 destruct - |#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)). - -lemma min_input_def : ∀h,i,x. - min_input h i x = min (x -i) (S i) (λy.termb i y (h (S i) y)). -// qed. - -lemma min_input_i: ∀h,i,x. x ≤ i → min_input h i x = S i. -#h #i #x #lexi >min_input_def -cut (x - i = 0) [@sym_eq /2 by eq_minus_O/] #Hcut // -qed. - -lemma min_input_to_terminate: ∀h,i,x. - min_input h i x = x → {i ⊙ x} ↓ (h (S i) x). -#h #i #x #Hminx -cases (decidable_le (S i) x) #Hix - [cases (true_or_false (termb i x (h (S i) x))) #Hcase - [@termb_true_to_term // - |min_input_def in Hminx; #Hminx >Hminx in ⊢ (%→?); - min_input_i in Hminx; - [#eqix >eqix in Hix; * /2/ | @le_S_S_to_le @not_le_to_lt //] - ] -qed. - -lemma min_input_to_lt: ∀h,i,x. - min_input h i x = x → i < x. -#h #i #x #Hminx cases (decidable_le (S i) x) // -#ltxi @False_ind >min_input_i in Hminx; - [#eqix >eqix in ltxi; * /2/ | @le_S_S_to_le @not_le_to_lt //] -qed. - -lemma le_to_min_input: ∀h,i,x,x1. x ≤ x1 → - min_input h i x = x → min_input h i x1 = x. -#h #i #x #x1 #lex #Hminx @(min_exists … (le_S_S … lex)) - [@(fmin_true … (sym_eq … Hminx)) // - |@(min_input_to_lt … Hminx) - |#j #H1 g_def cut (x-u = 0) [/2 by minus_le_minus_minus_comm/] -#eq0 >eq0 normalize // qed. - -lemma g_lt : ∀h,i,x. min_input h i x = x → - out i x (h (S i) x) < g h 0 x. -#h #i #x #H @le_S_S @(le_MaxI … i) /2 by min_input_to_lt/ -qed. - -lemma max_neq0 : ∀a,b. max a b ≠ 0 → a ≠ 0 ∨ b ≠ 0. -#a #b whd in match (max a b); cases (true_or_false (leb a b)) #Hcase >Hcase - [#H %2 @H | #H %1 @H] -qed. - -definition almost_equal ≝ λf,g:nat → nat. ¬ ∀nu.∃x. nu < x ∧ f x ≠ g x. -interpretation "almost equal" 'napart f g = (almost_equal f g). - -lemma eventually_cancelled: ∀h,u.¬∀nu.∃x. nu < x ∧ - max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) ≠ 0. -#h #u elim u - [normalize % #H cases (H u) #x * #_ * #H1 @H1 // - |#u0 @not_to_not #Hind #nu cases (Hind nu) #x * #ltx - cases (true_or_false (eqb (min_input h (u0+O) x) x)) #Hcase - [>bigop_Strue [2:@Hcase] #Hmax cases (max_neq0 … Hmax) -Hmax - [2: #H %{x} % // bigop_Sfalse - [#H %{x1} % [@transitive_lt //| (le_to_min_input … (eqb_true_to_eq … Hcase)) - [@lt_to_not_eq @ltx1 | @lt_to_le @ltx1] - ] - |>bigop_Sfalse [2:@Hcase] #H %{x} % // (bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA) - [>H // |@lt_to_le @(le_to_lt_to_lt …ltx) /2 by le_maxr/ |//] -qed. - -(******************************** Condition 2 *********************************) -definition total ≝ λf.λx:nat. Some nat (f x). - -lemma exists_to_exists_min: ∀h,i. (∃x. i < x ∧ {i ⊙ x} ↓ h (S i) x) → ∃y. min_input h i y = y. -#h #i * #x * #ltix #Hx %{(min_input h i x)} @min_spec_to_min @found // - [@(f_min_true (λy:ℕ.termb i y (h (S i) y))) %{x} % [% // | @term_to_termb_true //] - |#y #leiy #lty @(lt_min_to_false ????? lty) // - ] -qed. - -lemma condition_2: ∀h,i. code_for (total (g h 0)) i → ¬∃x. itermy >Hr -#H @(absurd ? H) @le_to_not_lt @le_n -qed. - - -(********************************* complexity *********************************) - -(* We assume operations have a minimal structural complexity MSC. -For instance, for time complexity, MSC is equal to the size of input. -For space complexity, MSC is typically 0, since we only measure the -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) *) - -definition C ≝ λs,i.∃c.∃a.∀x.a ≤ x → ∃y. - U i x (c*(s x)) = Some ? 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 ??); associative_times @le_times // @Ha1 @(transitive_le … lean) // - ] -qed. - -lemma timesc_CF: ∀s,f,c.CF (λx.c*s x) f → CF s f. -#s #f #c @O_to_CF @O_times_c -qed. - -(********************************* composition ********************************) -axiom CF_comp: ∀f,g,sf,sg,sh. CF sg g → CF sf f → - O sh (λx. sg x + sf (g x)) → CF sh (f ∘ g). - -lemma CF_comp_ext: ∀f,g,h,sh,sf,sg. CF sg g → CF sf f → - (∀x.f(g x) = h x) → O sh (λx. sg x + sf (g x)) → CF sh h. -#f #g #h #sh #sf #sg #Hg #Hf #Heq #H @(ext_CF (f ∘ g)) - [#n normalize @Heq | @(CF_comp … H) //] -qed. - -(* primitve recursion *) - -let rec prim_rec (k,h:nat →nat) n m on n ≝ - match n with - [ O ⇒ k m - | S a ⇒ h 〈a,〈prim_rec k h a m, m〉〉]. - -lemma prim_rec_S: ∀k,h,n,m. - prim_rec k h (S n) m = h 〈n,〈prim_rec k h n m, m〉〉. -// qed. - -definition unary_pr ≝ λk,h,x. prim_rec k h (fst x) (snd x). - -let rec prim_rec_compl (k,h,sk,sh:nat →nat) n m on n ≝ - match n with - [ O ⇒ sk m - | S a ⇒ prim_rec_compl k h sk sh a m + sh (prim_rec k h a m)]. - -axiom CF_prim_rec: ∀k,h,sk,sh,sf. CF sk k → CF sh h → - O sf (unary_pr sk (λx. fst (snd x) + sh 〈fst x,〈unary_pr k h 〈fst x,snd (snd x)〉,snd (snd x)〉〉)) - → CF sf (unary_pr k h). - -(* falso ???? -lemma prim_rec_O: ∀k1,h1,k2,h2. O k1 k2 → O h1 h2 → - O (unary_pr k1 h1) (unary_pr k2 h2). -#k1 #h1 #k2 #h2 #HO1 #HO2 whd *) - - -(**************************** primitive operations*****************************) - -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_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 ***********************************) - -axiom CF_eqb: ∀h,f,g. - CF h f → CF h g → CF h (λx.eqb (f x) (g x)). - -(*********************************** maximum **********************************) - -axiom CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s. - CF ha a → CF hb b → CF hp p → CF hf f → - O s (λx.ha x + hb x + ∑_{i ∈[a x ,b x[ }(hp 〈i,x〉 + hf 〈i,x〉)) → - CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). - -(******************************** minimization ********************************) - -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〉)). - -(************************************* smn ************************************) -axiom smn: ∀f,s. CF s f → ∀x. CF (λy.s 〈x,y〉) (λy.f 〈x,y〉). - -(****************************** constructibility ******************************) - -definition constructible ≝ λs. CF s s. - -lemma constr_comp : ∀s1,s2. constructible s1 → constructible s2 → - (∀x. x ≤ s2 x) → constructible (s2 ∘ s1). -#s1 #s2 #Hs1 #Hs2 #Hle @(CF_comp … Hs1 Hs2) @O_plus @le_to_O #x [@Hle | //] -qed. - -lemma ext_constr: ∀s1,s2. (∀x.s1 x = s2 x) → - constructible s1 → constructible s2. -#s1 #s2 #Hext #Hs1 @(ext_CF … Hext) @(monotonic_CF … Hs1) #x >Hext // -qed. - -lemma constr_prim_rec: ∀s1,s2. constructible s1 → constructible s2 → - (∀n,r,m. 2 * r ≤ s2 〈n,〈r,m〉〉) → constructible (unary_pr s1 s2). -#s1 #s2 #Hs1 #Hs2 #Hincr @(CF_prim_rec … Hs1 Hs2) whd %{2} %{0} -#x #_ lapply (surj_pair x) * #a * #b #eqx >eqx whd in match (unary_pr ???); ->fst_pair >snd_pair -whd in match (unary_pr ???); >fst_pair >snd_pair elim a - [normalize // - |#n #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair - >prim_rec_S @transitive_le [| @(monotonic_le_plus_l … Hind)] - @transitive_le [| @(monotonic_le_plus_l … (Hincr n ? b))] - whd in match (unary_pr ???); >fst_pair >snd_pair // - ] -qed. - -(********************************* simulation *********************************) - -axiom sU : nat → nat. - -axiom monotonic_sU: ∀i1,i2,x1,x2,s1,s2. i1 ≤ i2 → x1 ≤ x2 → s1 ≤ s2 → - sU 〈i1,〈x1,s1〉〉 ≤ sU 〈i2,〈x2,s2〉〉. - -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) -#b1 * #c1 #eqy >eqy -eqy -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〉〉. - -definition pU_unary ≝ λp. pU (fst p) (fst (snd p)) (snd (snd p)). - -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. - - -(******************** complexity of g ********************) - -definition unary_g ≝ λh.λux. g h (fst ux) (snd ux). -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 (auxg h) → CF s (unary_g h). -#h #s #H1 @(CF_compS ? (auxg h) H1) -qed. - -definition aux1g ≝ - λh,ux. max_{i ∈[fst ux,snd ux[ | (λp. eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) 〈i,ux〉} - ((λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) 〈i,ux〉). - -lemma eq_aux : ∀h,x.aux1g h x = auxg h x. -#h #x @same_bigop - [#n #_ >fst_pair >snd_pair // |#n #_ #_ >fst_pair >snd_pair //] -qed. - -lemma compl_g2 : ∀h,s1,s2,s. - CF s1 - (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) → - CF s2 - (λ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 (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 (λ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 min_input_def whd in ⊢ (??%?); >minus_S_S @min_f_g #i #_ #_ -whd in ⊢ (??%%); >fst_pair >snd_pair // -qed. - -definition termb_aux ≝ λh. - termb_unary ∘ λp.〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉. - -lemma compl_g4 : ∀h,s1,s. - (CF s1 - (λ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 (λ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 - |@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //] -qed. - -(************************* a couple of technical lemmas ***********************) -lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0. -#a elim a // #n #Hind * - [#H @False_ind /2 by absurd/ | #b normalize #H @Hind @le_S_S_to_le /2/] -qed. - -lemma sigma_bound: ∀h,a,b. monotonic nat le h → - ∑_{i ∈ [a,S b[ }(h i) ≤ (S b-a)*h b. -#h #a #b #H cases (decidable_le a b) - [#leab cut (b = pred (S b - a + a)) - [Hb in match (h b); - generalize in match (S b -a); - #n elim n - [// - |#m #Hind >bigop_Strue [2://] @le_plus - [@H @le_n |@(transitive_le … Hind) @le_times [//] @H //] - ] - |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba - cut (S b -a = 0) [@minus_to_0 //] #Hcut >Hcut // - ] -qed. - -lemma sigma_bound_decr: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) → - ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a. -#h #a #b #H cases (decidable_le a b) - [#leab cut ((b -a) +a ≤ b) [/2 by le_minus_to_plus_r/] generalize in match (b -a); - #n elim n - [// - |#m #Hind >bigop_Strue [2://] #Hm - cut (m+a ≤ b) [@(transitive_le … Hm) //] #Hm1 - @le_plus [@H // |@(transitive_le … (Hind Hm1)) //] - ] - |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba - cut (b -a = 0) [@minus_to_0 @lt_to_le @ltba] #Hcut >Hcut // - ] -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〉)). -#s1 #Hs1 %{1} %{0} #n #_ >commutative_times minus_S_S //] -qed. - -lemma coroll2: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) → -O (λx.(snd x - fst x)*s1 〈fst x,x〉) (λx.∑_{i ∈[fst x,snd x[ }(s1 〈i,x〉)). -#s1 #Hs1 %{1} %{0} #n #_ >commutative_times fst_pair >snd_pair >fst_pair >snd_pair // ] -@(CF_comp … (λx.MSC 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 - [@O_plus - [@(O_trans … (λx.MSC (fst x) + MSC (max (fst (snd x)) (snd (snd x))))) - [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx - >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) - >distributive_times_plus @le_plus [//] - cases (surj_pair b) #c * #d #eqb >eqb - >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) - whd in ⊢ (??%); @le_plus - [@monotonic_MSC @(le_maxl … (le_n …)) - |>commutative_times fst_pair >snd_pair // qed. - -lemma le_big : ∀x. x ≤ big x. -#x cases (surj_pair x) #a * #b #eqx >eqx @le_pair >fst_pair >snd_pair -[@(le_maxl … (le_n …)) | @(le_maxr … (le_n …))] -qed. - -definition faux2 ≝ λh. - (λx.MSC x + (snd (snd x)-fst x)* - (λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉). - -lemma compl_g7: ∀h. - constructible (λx. h (fst x) (snd x)) → - (∀n. monotonic ? le (h n)) → - CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉) - (λp:ℕ.min_input h (fst p) (snd (snd p))). -#h #hcostr #hmono @(monotonic_CF … (faux2 h)) - [#n normalize >fst_pair >snd_pair //] -@compl_g5 [2:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair ->fst_pair >snd_pair @monotonic_sU // @hmono @lexy -qed. - -lemma compl_g71: ∀h. - 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))〉〉) - (λp:ℕ.min_input h (fst p) (snd (snd p))). -#h #hcostr #hmono @(monotonic_CF … (compl_g7 h hcostr hmono)) #x -@le_plus [@monotonic_MSC //] -cases (decidable_le (fst x) (snd(snd x))) - [#Hle @le_times // @monotonic_sU - |#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt] - ] -qed. - -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 (MSC x)) >commutative_times @le_times - [// | @monotonic_MSC // ]] -@(O_trans … (coroll2 ??)) - [#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 [//|@(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 - [/2 by monotonic_le_minus_r/ - |@monotonic_sU // @hantimono [@le_S_S // |@ltb] - ] - |@monotonic_sU // @hantimono [@le_S_S // |@ltb] - ] - |@le_to_O #n >fst_pair >snd_pair - cut (max (fst n) n = n) [normalize >le_to_leb_true //] #Hmax >Hmax - >associative_plus >distributive_times_plus - @le_plus [@le_times [@le_S // |>big_def >Hmax //] |//] - ] -qed. - -definition sg ≝ λh,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)〉〉. - -lemma sg_def : ∀h,a,b. - sg h 〈a,b〉 = (S (b-a))*MSC 〈〈a,b〉,〈a,b〉〉 + - (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉. -#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair // -qed. - -lemma compl_g11 : ∀h. - 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) (unary_g h). -#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham) -qed. - -(**************************** closing the argument ****************************) - -let rec h_of_aux (r:nat →nat) (c,d,b:nat) on d : nat ≝ - match d with - [ O ⇒ c - | S d1 ⇒ (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + - d*(S d)*sU 〈〈b-d,b〉,〈b,r (h_of_aux r c d1 b)〉〉]. - -lemma h_of_aux_O: ∀r,c,b. - h_of_aux r c O b = c. -// qed. - -lemma h_of_aux_S : ∀r,c,d,b. - h_of_aux r c (S d) b = - (S (S d))*(MSC 〈〈b-(S d),b〉,〈b-(S d),b〉〉) + - (S d)*(S (S d))*sU 〈〈b-(S d),b〉,〈b,r(h_of_aux r c d b)〉〉. -// qed. - -lemma h_of_aux_prim_rec : ∀r,c,n,b. h_of_aux r c n b = - prim_rec (λx.c) - (λx.let d ≝ S(fst x) in - let b ≝ snd (snd x) in - (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + - d*(S d)*sU 〈〈b-d,b〉,〈b,r (fst (snd x))〉〉) n b. -#r #c #n #b elim n - [>h_of_aux_O normalize // - |#n1 #Hind >h_of_aux_S >prim_rec_S >snd_pair >snd_pair >fst_pair - >fst_pair fst_pair >snd_pair >(minus_to_0 … Hle) // -qed. - -lemma h_of_def: ∀r,a,b.h_of r 〈a,b〉 = - let m ≝ max a b in - h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (b - a) b. -#r #a #b normalize >fst_pair >snd_pair // -qed. - -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. -#r #Hr #monor #d #d1 lapply d -d elim d1 - [#d #c #c1 #b #b1 #Hc #Hd @(le_n_O_elim ? Hd) #leb - >h_of_aux_O >h_of_aux_O // - |#m #Hind #d #c #c1 #b #b1 #lec #led #leb cases (le_to_or_lt_eq … led) - [#ltd @(transitive_le … (Hind … lec ? leb)) [@le_S_S_to_le @ltd] - >h_of_aux_S @(transitive_le ???? (le_plus_n …)) - >(times_n_1 (h_of_aux r c1 m b1)) in ⊢ (?%?); - >commutative_times @le_times [//|@(transitive_le … (Hr ?)) @sU_le] - |#Hd >Hd >h_of_aux_S >h_of_aux_S - cut (b-S m ≤ b1 - S m) [/2 by monotonic_le_minus_l/] #Hb1 - @le_plus [@le_times //] - [@monotonic_MSC @le_pair @le_pair // - |@le_times [//] @monotonic_sU - [@le_pair // |// |@monor @Hind //] - ] - ] - ] -qed. - -lemma mono_h_of2: ∀r.(∀x. x ≤ r x) → monotonic ? le r → - ∀i,b,b1. b ≤ b1 → h_of r 〈i,b〉 ≤ h_of r 〈i,b1〉. -#r #Hr #Hmono #i #a #b #leab >h_of_def >h_of_def -cut (max i a ≤ max i b) - [@to_max - [@(le_maxl … (le_n …))|@(transitive_le … leab) @(le_maxr … (le_n …))]] -#Hmax @(mono_h_of_aux r Hr Hmono) - [@monotonic_MSC @le_pair @le_pair @Hmax |/2 by monotonic_le_minus_l/ |@leab] -qed. - -axiom h_of_constr : ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - constructible (h_of r). - -lemma speed_compl: ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))). -#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …)) - [#x cases (surj_pair x) #a * #b #eqx >eqx - >sg_def cases (decidable_le b a) - [#leba >(minus_to_0 … leba) normalize in ⊢ (?%?); - h_of_def - cut (max a b = a) - [normalize cases (le_to_or_lt_eq … leba) - [#ltba >(lt_to_leb_false … ltba) % - |#eqba (le_to_leb_true … (le_n ?)) % ]] - #Hmax >Hmax normalize >(minus_to_0 … leba) normalize - @monotonic_MSC @le_pair @le_pair // - |#ltab >h_of_def >h_of_def - cut (max a b = b) - [normalize >(le_to_leb_true … ) [%] @lt_to_le @not_le_to_lt @ltab] - #Hmax >Hmax - 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) - [%{(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 - [@lt_to_le @not_le_to_lt // | //]] #eqd2 >eqd2 - normalize // - ] - |#n #a #b #leab #lebn >h_of_def >h_of_def - cut (max a n = n) - [normalize >le_to_leb_true [%|@(transitive_le … leab lebn)]] #Hmaxa - cut (max b n = n) - [normalize >(le_to_leb_true … lebn) %] #Hmaxb - >Hmaxa >Hmaxb @Hmono @(mono_h_of_aux r … Hr Hmono) // /2 by monotonic_le_minus_r/ - |#n #a #b #leab @Hmono @(mono_h_of2 … Hr Hmono … leab) - |@(constr_comp … Hconstr Hr) @(ext_constr (h_of r)) - [#x cases (surj_pair x) #a * #b #eqx >eqx >fst_pair >snd_pair //] - @(h_of_constr r Hr Hmono Hconstr) - ] -qed. - -lemma speed_compl_i: ∀r:nat →nat. - (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∀i. CF (λx.h_of r 〈i,x〉) (λx.g (λi,x. r(h_of r 〈i,x〉)) i x). -#r #Hr #Hmono #Hconstr #i -@(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 Hconstr)) #n // -qed. - -(**************************** the speedup theorem *****************************) -theorem pseudo_speedup: - ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∃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 #Hconstr -(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * -#Hcodei #HCi -(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} -(* sg is (λx.h_of r 〈i,x〉) *) -%{(λx. h_of r 〈S i,x〉)} -lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg -%[%[@condition_1 |@Hg] - |cases Hg #H1 * #j * #Hcodej #HCj - lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) - cases HCi #m * #a #Ha %{m} %{(max (S i) a)} #n #ltin @lt_to_le @not_le_to_lt - @(not_to_not … Hcond2) -Hcond2 #Hlesf %{n} % - [@(transitive_le … ltin) @(le_maxl … (le_n …))] - cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] - #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) // - ] -qed. - -theorem pseudo_speedup': - ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → - ∃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 #Hconstr -(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * -#Hcodei #HCi -(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) -%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} -(* sg is (λx.h_of r 〈i,x〉) *) -%{(λx. h_of r 〈S i,x〉)} -lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg -%[%[@condition_1 |@Hg] - |cases Hg #H1 * #j * #Hcodej #HCj - lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) - cases HCi #m * #a #Ha - %{m} %{(max (S i) a)} #n #ltin @not_le_to_lt @(not_to_not … Hcond2) -Hcond2 #Hlesf - %{n} % [@(transitive_le … ltin) @(le_maxl … (le_n …))] - cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] - #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) - @Hmono @(mono_h_of2 … Hr Hmono … ltin) - ] -qed. - diff --git a/matita/matita/lib/reverse_complexity/almost.ma b/matita/matita/lib/reverse_complexity/almost.ma new file mode 100644 index 000000000..75578d0f9 --- /dev/null +++ b/matita/matita/lib/reverse_complexity/almost.ma @@ -0,0 +1,24 @@ +include "reverse_complexity/speedup.ma". + +definition aes : (nat → nat) → (nat → nat) → Prop ≝ λf,g. + ∃n.∀m. n ≤ m → f m = g m. + +lemma CF_almost: ∀f,g,s. CF s g → aes g f → CF s f. +#f #g #s #CFg * #n lapply CFg -CFg lapply g -g +elim n + [#g #CFg #H @(ext_CF … g) [#m @H // |//] + |#i #Hind #g #CFg #H + @(Hind (λx. if eqb i x then f i else g x)) + [@CF_if + [@(O_to_CF … MSC) [@le_to_O @(MSC_in … CFg)] @CF_eqb // + |@(O_to_CF … MSC) [@le_to_O @(MSC_in … CFg)] @CF_const + |@CFg + ] + |#m #leim cases (le_to_or_lt_eq … leim) + [#ltim lapply (lt_to_not_eq … ltim) #noteqim + >(not_eq_to_eqb_false … noteqim) @H @ltim + |#eqim >eqim >eqb_n_n // + ] + ] + ] +qed. \ No newline at end of file diff --git a/matita/matita/lib/reverse_complexity/basics.ma b/matita/matita/lib/reverse_complexity/basics.ma new file mode 100644 index 000000000..de383d1d0 --- /dev/null +++ b/matita/matita/lib/reverse_complexity/basics.ma @@ -0,0 +1,106 @@ +include "basics/types.ma". +include "arithmetics/nat.ma". + +(********************************** pairing ***********************************) +axiom pair: nat → nat → nat. +axiom fst : nat → nat. +axiom snd : nat → nat. + +interpretation "abstract pair" 'pair f g = (pair f g). + +axiom fst_pair: ∀a,b. fst 〈a,b〉 = a. +axiom snd_pair: ∀a,b. snd 〈a,b〉 = b. +axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉. + +axiom le_fst : ∀p. fst p ≤ p. +axiom le_snd : ∀p. snd p ≤ p. +axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉. + +lemma expand_pair: ∀x. x = 〈fst x, snd x〉. +#x lapply (surj_pair x) * #a * #b #Hx >Hx >fst_pair >snd_pair // +qed. + +(************************************* U **************************************) +axiom U: nat → nat →nat → option nat. + +axiom monotonic_U: ∀i,x,n,m,y.n ≤m → + U i x n = Some ? y → U i x m = Some ? y. + +lemma unique_U: ∀i,x,n,m,yn,ym. + U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. +#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. + +notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. + +lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n. +#i #x #n normalize cases (U i x n) + [%2 % * #y #H destruct|#y %1 %{y} //] +qed. + +lemma monotonic_terminate: ∀i,x,n,m. + n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m. +#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) // +qed. + +definition termb ≝ λi,x,t. + match U i x t with [None ⇒ false |Some y ⇒ true]. + +lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t. +#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] +qed. + +lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true. +#i #x #t * #y #H normalize >H // +qed. + +definition out ≝ λi,x,r. + match U i x r with [ None ⇒ 0 | Some z ⇒ z]. + +definition bool_to_nat: bool → nat ≝ + λb. match b with [true ⇒ 1 | false ⇒ 0]. + +coercion bool_to_nat. + +definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉. + +lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y. +#i #x #r #y % normalize + [cases (U i x r) normalize + [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H] + #H1 destruct + |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1] + #H1 // + ] + |#H >H //] +qed. + +lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?. +#i #x #r % normalize + [cases (U i x r) normalize // + #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1] + #H1 destruct + |#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. + + +definition total ≝ λf.λx:nat. Some nat (f x). + + diff --git a/matita/matita/lib/reverse_complexity/big_O.ma b/matita/matita/lib/reverse_complexity/big_O.ma new file mode 100644 index 000000000..833572ea2 --- /dev/null +++ b/matita/matita/lib/reverse_complexity/big_O.ma @@ -0,0 +1,89 @@ + +include "arithmetics/nat.ma". +include "basics/sets.ma". + +(******************************** 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. + +lemma le_to_O: ∀s1,s2. (∀x.s1 x ≤ s2 x) → O s2 s1. +#s1 #s2 #Hle %{1} %{0} #n #_ 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_plus_r: ∀f,s1,s2. O s2 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_times_c: ∀f,c. O f (λx:ℕ.c*f x). +#f #c %{c} %{0} // +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 (times_n_1 (s n0)) in ⊢ (?%?); >commutative_times @H // +qed. + +lemma o_trans: ∀s1,s2,s3. o s2 s1 → o s3 s2 → o s3 s1. +#s1 #s2 #s3 #H1 #H2 #c cases (H1 c) #n1 -H1 #H1 cases (H2 1) #n2 -H2 #H2 +%{(max n1 n2)} #n #Hmax +@(transitive_lt … (H1 ??)) [@(le_maxl … Hmax)] +>(times_n_1 (s2 n)) in ⊢ (?%?); >commutative_times @H2 @(le_maxr … Hmax) +qed. diff --git a/matita/matita/lib/reverse_complexity/bigops_compl.ma b/matita/matita/lib/reverse_complexity/bigops_compl.ma new file mode 100644 index 000000000..032174a0c --- /dev/null +++ b/matita/matita/lib/reverse_complexity/bigops_compl.ma @@ -0,0 +1,607 @@ +include "reverse_complexity/complexity.ma". +include "reverse_complexity/sigma_diseq.ma". + +include alias "reverse_complexity/basics.ma". + +lemma bigop_prim_rec: ∀a,b,c,p,f,x. + bigop (b x-a x) (λi:ℕ.p 〈i+a x,x〉) ? (c 〈a x,x〉) plus (λi:ℕ.f 〈i+a x,x〉) = + prim_rec c + (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉 + then plus (f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉) (fst (snd i)) + else fst (snd i)) (b x-a x) 〈a x ,x〉. +#a #b #c #p #f #x normalize elim (b x-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma bigop_prim_rec_dec: ∀a,b,c,p,f,x. + bigop (b x-a x) (λi:ℕ.p 〈b x -i,x〉) ? (c 〈b x,x〉) plus (λi:ℕ.f 〈b x-i,x〉) = + prim_rec c + (λi.if p 〈fst (snd (snd i)) - fst i,snd (snd (snd i))〉 + then plus (f 〈fst (snd (snd i)) - fst i,snd (snd (snd i))〉) (fst (snd i)) + else fst (snd i)) (b x-a x) 〈b x ,x〉. +#a #b #c #p #f #x normalize elim (b x-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + cases (true_or_false (p 〈b x - i,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma bigop_prim_rec_dec1: ∀a,b,c,p,f,x. + bigop (S(b x)-a x) (λi:ℕ.p 〈b x - i,x〉) ? (c 〈b x,x〉) plus (λi:ℕ.f 〈b x- i,x〉) = + prim_rec c + (λi.if p 〈fst (snd (snd i)) - (fst i),snd (snd (snd i))〉 + then plus (f 〈fst (snd (snd i)) - (fst i),snd (snd (snd i))〉) (fst (snd i)) + else fst (snd i)) (S(b x)-a x) 〈b x,x〉. +#a #b #c #p #f #x elim (S(b x)-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + cases (true_or_false (p 〈b x - i,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma sum_prim_rec1: ∀a,b,p,f,x. + ∑_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) = + prim_rec (λi.0) + (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉 + then f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉 + fst (snd i) + else fst (snd i)) (b x-a x) 〈a x ,x〉. +#a #b #p #f #x elim (b x-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma bigop_plus_c: ∀k,p,f,c. + c k + bigop k (λi.p i) ? 0 plus (λi.f i) = + bigop k (λi.p i) ? (c k) plus (λi.f i). +#k #p #f elim k [normalize //] +#i #Hind #c cases (true_or_false (p i)) #Hcase +[>bigop_Strue // >bigop_Strue // (commutative_plus ? (f i)) + >associative_plus @eq_f @Hind +|>bigop_Sfalse // >bigop_Sfalse // +] +qed. + + +(*********************************** maximum **********************************) + +lemma max_gen: ∀a,b,p,f,x. a ≤b → + max_{i ∈[a,b[ | p 〈i,x〉 }(f 〈i,x〉) = max_{i < b | leb a i ∧ p 〈i,x〉 }(f 〈i,x〉). +#a #b #p #f #x @(bigop_I_gen ????? MaxA) +qed. + +lemma max_prim_rec_base: ∀a,b,p,f,x. a ≤b → + max_{i < b| p 〈i,x〉 }(f 〈i,x〉) = + prim_rec (λi.0) + (λi.if p 〈fst i,x〉 then max (f 〈fst i,snd (snd i)〉) (fst (snd i)) else fst (snd i)) b x. +#a #b #p #f #x #leab >max_gen // elim b + [normalize // + |#i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair + cases (true_or_false (p 〈i,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma max_prim_rec: ∀a,b,p,f,x. a ≤b → + max_{i ∈[a,b[ | p 〈i,x〉 }(f 〈i,x〉) = + prim_rec (λi.0) + (λi.if leb a (fst i) ∧ p 〈fst i,x〉 then max (f 〈fst i,snd (snd i)〉) (fst (snd i)) else fst (snd i)) b x. +#a #b #p #f #x #leab >max_gen // elim b + [normalize // + |#i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair + cases (true_or_false (leb a i ∧ p 〈i,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma max_prim_rec1: ∀a,b,p,f,x. + max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) = + prim_rec (λi.0) + (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉 + then max (f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉) (fst (snd i)) + else fst (snd i)) (b x-a x) 〈a x ,x〉. +#a #b #p #f #x elim (b x-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +(* the argument is 〈b-a,〈a,x〉〉 *) + +definition max_unary_pr ≝ λp,f.unary_pr (λx.0) + (λi. + let k ≝ fst i in + let r ≝ fst (snd i) in + let a ≝ fst (snd (snd i)) in + let x ≝ snd (snd (snd i)) in + if p 〈k + a,x〉 then max (f 〈k+a,x〉) r else r). + +lemma max_unary_pr1: ∀a,b,p,f,x. + max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) = + ((max_unary_pr p f) ∘ (λx.〈b x - a x,〈a x,x〉〉)) x. +#a #b #p #f #x normalize >fst_pair >snd_pair @max_prim_rec1 +qed. + +definition aux_compl ≝ λhp,hf.λi. + let k ≝ fst i in  + let r ≝ fst (snd i) in  + let a ≝ fst (snd (snd i)) in  + let x ≝ snd (snd (snd i)) in  + hp 〈k+a,x〉 + hf 〈k+a,x〉 + (* was MSC r*) MSC i . + +definition aux_compl1 ≝ λhp,hf.λi. + let k ≝ fst i in  + let r ≝ fst (snd i) in  + let a ≝ fst (snd (snd i)) in  + let x ≝ snd (snd (snd i)) in  + hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r. + +lemma aux_compl1_def: ∀k,r,m,hp,hf. + aux_compl1 hp hf 〈k,〈r,m〉〉 = + let a ≝ fst m in  + let x ≝ snd m in  + hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r. +#k #r #m #hp #hf normalize >fst_pair >snd_pair >snd_pair >fst_pair // +qed. + +lemma aux_compl1_def1: ∀k,r,a,x,hp,hf. + aux_compl1 hp hf 〈k,〈r,〈a,x〉〉〉 = hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r. +#k #r #a #x #hp #hf normalize >fst_pair >snd_pair >snd_pair >fst_pair +>fst_pair >snd_pair // +qed. + +axiom Oaux_compl: ∀hp,hf. O (aux_compl1 hp hf) (aux_compl hp hf). + +lemma MSC_max: ∀f,h,x. CF h f → MSC (max_{i < x}(f i)) ≤ max_{i < x}(h i). +#f #h #x #hf elim x // #i #Hind >bigop_Strue [|//] >bigop_Strue [|//] +whd in match (max ??); +cases (true_or_false (leb (f i) (bigop i (λi0:ℕ.true) ? 0 max(λi0:ℕ.f i0)))) +#Hcase >Hcase + [@(transitive_le … Hind) @le_maxr // + |@(transitive_le … (MSC_out … hf i)) @le_maxl // + ] +qed. + +lemma CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s. + CF ha a → CF hb b → CF hp p → CF hf f → + O s (λx.ha x + hb x + + (∑_{i ∈[a x ,b x[ } + (hp 〈i,x〉 + hf 〈i,x〉 + max_{i ∈ [a x, b x [ }(hf 〈i,x〉)))) → + CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). +#a #b #p #f #ha #hb #hp #hf #s #CFa #CFb #CFp #CFf #HO +@ext_CF1 [|#x @max_unary_pr1] +@(CF_comp ??? (λx.ha x + hb x)) + [|@CF_comp_pair + [@CF_minus [@(O_to_CF … CFb) @O_plus_r // |@(O_to_CF … CFa) @O_plus_l //] + |@CF_comp_pair + [@(O_to_CF … CFa) @O_plus_l // + | @(O_to_CF … CF_id) @O_plus_r @le_to_O @(MSC_in … CFb) + ] + ] + |@(CF_prim_rec … MSC … (aux_compl1 hp hf)) + [@CF_const + |@(O_to_CF … (Oaux_compl … )) + @CF_if + [@(CF_comp p … MSC … CFp) + [@CF_comp_pair + [@CF_plus [@CF_fst| @CF_comp_fst @CF_comp_snd @CF_snd] + |@CF_comp_snd @CF_comp_snd @CF_snd] + |@le_to_O #x normalize >commutative_plus >associative_plus @le_plus // + ] + |@CF_max1 + [@(CF_comp f … MSC … CFf) + [@CF_comp_pair + [@CF_plus [@CF_fst| @CF_comp_fst @CF_comp_snd @CF_snd] + |@CF_comp_snd @CF_comp_snd @CF_snd] + |@le_to_O #x normalize >commutative_plus // + ] + |@CF_comp_fst @(monotonic_CF … CF_snd) normalize // + ] + |@CF_comp_fst @(monotonic_CF … CF_snd) normalize // + ] + |@O_refl + ] + |@(O_trans … HO) -HO + @(O_trans ? (λx:ℕ + .ha x+hb x + +bigop (b x-a x) (λi:ℕ.true) ? (MSC 〈a x,x〉) plus + (λi:ℕ + .(λi0:ℕ + .hp 〈i0,x〉+hf 〈i0,x〉 + +bigop (b x-a x) (λi1:ℕ.true) ? 0 max + (λi1:ℕ.(λi2:ℕ.hf 〈i2,x〉) (i1+a x))) (i+a x)))) + [ + @le_to_O #n @le_plus // whd in match (unary_pr ???); + >fst_pair >snd_pair >bigop_prim_rec elim (b n - a n) + [normalize // + |#x #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >aux_compl1_def1 + >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >fst_pair + >snd_pair normalize in ⊢ (??%); >commutative_plus @le_plus + [-Hind @le_plus // normalize >fst_pair >snd_pair + @(transitive_le ? (bigop x (λi1:ℕ.true) ? 0 (λn0:ℕ.λm:ℕ.if leb n0 m then m else n0 ) + (λi1:ℕ.hf 〈i1+a n,n〉))) + [elim x [normalize @MSC_le] + #x0 #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >snd_pair + >fst_pair >fst_pair cases (p 〈x0+a n,n〉) normalize + [cases (true_or_false (leb (f 〈x0+a n,n〉) + (prim_rec (λx00:ℕ.O) + (λi:ℕ + .if p 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉  + then if leb (f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉) + (fst (snd i))  + then fst (snd i)  + else f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉   + else fst (snd i) ) x0 〈a n,n〉))) #Hcase >Hcase normalize + [@(transitive_le … Hind) -Hind @(le_maxr … (le_n …)) + |@(transitive_le … (MSC_out … CFf …)) @(le_maxl … (le_n …)) + ] + |@(transitive_le … Hind) -Hind @(le_maxr … (le_n …)) + ] + |@(le_maxr … (le_n …)) + ] + |@(transitive_le … Hind) -Hind + generalize in match (bigop x (λi:ℕ.true) ? 0 max + (λi1:ℕ.(λi2:ℕ.hf 〈i2,n〉) (i1+a n))); #c + generalize in match (hf 〈x+a n,n〉); #c1 + elim x [//] #x0 #Hind + >prim_rec_S >prim_rec_S normalize >fst_pair >fst_pair >snd_pair + >snd_pair >snd_pair >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair + >fst_pair @le_plus + [@le_plus // cases (true_or_false (leb c1 c)) #Hcase + >Hcase normalize // @lt_to_le @not_le_to_lt @(leb_false_to_not_le ?? Hcase) + |@Hind + ] + ] + ] + |@O_plus [@O_plus_l //] @le_to_O #x + Hind +cases (true_or_false (eqb a (a+i))) #Hcase >Hcase normalize // +<(eqb_true_to_eq … Hcase) >H // +qed. + +lemma my_minim_nfa : ∀a,f,x,k. f〈a,x〉 = false → +my_minim a f x (S k) = my_minim (S a) f x k. +#a #f #x #k #H elim k + [normalize H >eq_to_eqb_true // + |#i #Hind >my_minim_S >Hind >my_minim_S Hcase + [>(my_minim_fa … Hcase) // | >Hind @sym_eq @(my_minim_nfa … Hcase) ] +qed. + +(* cambiare qui: togliere S *) + + +definition minim_unary_pr ≝ λf.unary_pr (λx.S(fst x)) + (λi. + let k ≝ fst i in + let r ≝ fst (snd i) in + let b ≝ fst (snd (snd i)) in + let x ≝ snd (snd (snd i)) in + if f 〈b-k,x〉 then b-k else r). + +definition compl_minim_unary ≝ λhf:nat → nat.λi. + let k ≝ fst i in + let b ≝ fst (snd (snd i)) in + let x ≝ snd (snd (snd i)) in + hf 〈b-k,x〉 + MSC 〈k,〈S b,x〉〉. + +definition compl_minim_unary_aux ≝ λhf,i. + let k ≝ fst i in + let r ≝ fst (snd i) in + let b ≝ fst (snd (snd i)) in + let x ≝ snd (snd (snd i)) in + hf 〈b-k,x〉 + MSC i. + +lemma compl_minim_unary_aux_def : ∀hf,k,r,b,x. + compl_minim_unary_aux hf 〈k,〈r,〈b,x〉〉〉 = hf 〈b-k,x〉 + MSC 〈k,〈r,〈b,x〉〉〉. +#hf #k #r #b #x normalize >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair // +qed. + +lemma compl_minim_unary_def : ∀hf,k,r,b,x. + compl_minim_unary hf 〈k,〈r,〈b,x〉〉〉 = hf 〈b-k,x〉 + MSC 〈k,〈S b,x〉〉. +#hf #k #r #b #x normalize >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair // +qed. + +lemma compl_minim_unary_aux_def2 : ∀hf,k,r,x. + compl_minim_unary_aux hf 〈k,〈r,x〉〉 = hf 〈fst x-k,snd x〉 + MSC 〈k,〈r,x〉〉. +#hf #k #r #x normalize >snd_pair >snd_pair >fst_pair // +qed. + +lemma compl_minim_unary_def2 : ∀hf,k,r,x. + compl_minim_unary hf 〈k,〈r,x〉〉 = hf 〈fst x-k,snd x〉 + MSC 〈k,〈S(fst x),snd x〉〉. +#hf #k #r #x normalize >snd_pair >snd_pair >fst_pair // +qed. + +lemma min_aux: ∀a,f,x,k. min (S k) (a x) (λi:ℕ.f 〈i,x〉) = + ((minim_unary_pr f) ∘ (λx.〈S k,〈a x + k,x〉〉)) x. +#a #f #x #k whd in ⊢ (???%); >fst_pair >snd_pair +lapply a -a (* @max_prim_rec1 *) +elim k + [normalize #a >fst_pair >snd_pair >fst_pair >snd_pair >snd_pair >fst_pair + fst_pair // + |#i #Hind #a normalize in ⊢ (??%?); >prim_rec_S >fst_pair >snd_pair + >fst_pair >snd_pair >snd_pair >fst_pair minus_S_S <(minus_plus_m_m (a x) i) // +qed. + +lemma minim_unary_pr1: ∀a,b,f,x. + μ_{i ∈[a x,b x]}(f 〈i,x〉) = + if leb (a x) (b x) then ((minim_unary_pr f) ∘ (λx.〈S (b x) - a x,〈b x,x〉〉)) x + else (a x). +#a #b #f #x cases (true_or_false (leb (a x) (b x))) #Hcase >Hcase + [cut (b x = a x + (b x - a x)) + [>commutative_plus minus_Sn_m [|@leb_true_to_le //] + >min_aux whd in ⊢ (??%?); eq_minus_O // @not_le_to_lt @leb_false_to_not_le // + ] +qed. + +axiom sum_inv: ∀a,b,f.∑_{i ∈ [a,S b[ }(f i) = ∑_{i ∈ [a,S b[ }(f (a + b - i)). + +(* +#a #b #f @(bigop_iso … plusAC) whd %{(λi.b - a - i)} %{(λi.b - a -i)} % + [%[#i #lti #_ normalize @eq_f >commutative_plus commutative_plus // + ] + |@(O_to_CF … MSC) + [@le_to_O #x normalize // + |@CF_minus + [@CF_comp_fst @CF_comp_snd @CF_snd |@CF_fst] + ] + |@CF_comp_fst @(monotonic_CF … CF_snd) normalize // + ] + |@O_plus + [@O_plus_l @O_refl + |@O_plus_r @O_ext2 [||#x >compl_minim_unary_aux_def2 @refl] + @O_trans [||@le_to_O #x >compl_minim_unary_def2 @le_n] + @O_plus [@O_plus_l //] + @O_plus_r + @O_trans [|@le_to_O #x @MSC_pair] @O_plus + [@le_to_O #x @monotonic_MSC @(transitive_le ???? (le_fst …)) + >fst_pair @le_n] + @O_trans [|@le_to_O #x @MSC_pair] @O_plus + [@le_to_O #x @monotonic_MSC @(transitive_le ???? (le_snd …)) + >snd_pair @(transitive_le ???? (le_fst …)) >fst_pair + normalize >snd_pair >fst_pair lapply (surj_pair x) + * #x1 * #x2 #Hx >Hx >fst_pair >snd_pair elim x1 // + #i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair + cases (f ?) [@le_S // | //]] + @le_to_O #x @monotonic_MSC @(transitive_le ???? (le_snd …)) >snd_pair + >(expand_pair (snd (snd x))) in ⊢ (?%?); @le_pair // + ] + ] + |cut (O s (λx.ha x + hb x + + ∑_{i ∈[a x ,S(b x)[ }(hf 〈a x+b x-i,x〉 + MSC 〈b x -(a x+b x-i),〈S(b x),x〉〉))) + [@(O_ext2 … HO) #x @eq_f @sum_inv] -HO #HO + @(O_trans … HO) -HO + @(O_trans ? (λx:ℕ.ha x+hb x + +bigop (S(b x)-a x) (λi:ℕ.true) ? + (MSC 〈b x,x〉) plus (λi:ℕ.(λj.hf j + MSC 〈b x - fst j,〈S(b (snd j)),snd j〉〉) 〈b x- i,x〉))) + [@le_to_O #n @le_plus // whd in match (unary_pr ???); + >fst_pair >snd_pair >(bigop_prim_rec_dec1 a b ? (λi.true)) + (* it is crucial to recall that the index is bound by S(b x) *) + cut (S (b n) - a n ≤ S (b n)) [//] + elim (S(b n) - a n) + [normalize // + |#x #Hind #lex >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair + >compl_minim_unary_def >prim_rec_S >fst_pair >snd_pair >fst_pair + >snd_pair >fst_pair >snd_pair >fst_pair whd in ⊢ (??%); >commutative_plus + @le_plus [2:@Hind @le_S @le_S_S_to_le @lex] -Hind >snd_pair + >minus_minus_associative // @le_S_S_to_le // + ] + |@O_plus [@O_plus_l //] @O_ext2 [||#x @bigop_plus_c] + @O_plus + [@O_plus_l @O_trans [|@le_to_O #x @MSC_pair] + @O_plus + [@O_plus_r @le_to_O @(MSC_out … CFb) + |@O_plus_r @le_to_O @(MSC_in … CFb) + ] + |@O_plus_r @(O_ext2 … (O_refl …)) #x @same_bigop + [//|#i #H #_ @eq_f2 [@eq_f @eq_f2 //|>fst_pair @eq_f @eq_f2 //] + ] + ] + ] + ] + |@(O_to_CF … CFa) @(O_trans … HO) @O_plus_l @O_plus_l @O_refl + ] + +qed. + +lemma CF_mu2: ∀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〉 + MSC〈S(b x),x〉)) → + CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)). +#a #b #f #sa #sb #sf #s #CFa #CFb #CFf #HO +@(O_to_CF … HO (CF_mu … CFa CFb CFf ?)) @O_plus [@O_plus_l @O_refl] +@O_plus_r @O_ext2 [|| #x @(bigop_op … plusAC)] +@O_plus [@le_to_O #x @le_sigma //] +@(O_trans ? (λx.∑_{i ∈[a x ,S(b x)[ }(MSC(b x -i)+MSC 〈S(b x),x〉))) + [@le_to_O #x @le_sigma //] +@O_ext2 [|| #x @(bigop_op … plusAC)] @O_plus + [@le_to_O #x @le_sigma // #i #lti #_ @(transitive_le … (MSC 〈S (b x),x〉)) // + @monotonic_MSC @(transitive_le … (S(b x))) // @le_S // + |@le_to_O #x @le_sigma // + ] +qed. + +(* uses MSC_S *) + +lemma CF_mu3: ∀a,b.∀f:nat →bool.∀sa,sb,sf,s. (∀x.sf x > 0) → + 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〉 + MSC〈b x,x〉)) → + CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)). +#a #b #f #sa #sb #sf #s #sfpos #CFa #CFb #CFf #HO +@(O_to_CF … HO (CF_mu2 … CFa CFb CFf ?)) @O_plus [@O_plus_l @O_refl] +@O_plus_r @O_ext2 [|| #x @(bigop_op … plusAC)] +@O_plus [@le_to_O #x @le_sigma //] +@le_to_O #x @le_sigma // #x #lti #_ >MSC_pair_eq >MSC_pair_eq 0) → + CF sa a → CF sb b → CF sf f → + O s (λx.sa x + sb x + (S(b x) - a x)*Max_{i ∈[a x ,S(b x)[ }(sf 〈i,x〉)) → + CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)). +#a #b #f #sa #sb #sf #s #sfpos #CFa #CFb #CFf #HO +@(O_to_CF … HO (CF_mu3 … sfpos CFa CFb CFf ?)) @O_plus [@O_plus_l @O_refl] +@O_ext2 [|| #x @(bigop_op … plusAC)] @O_plus_r @O_plus + [@le_to_O #x @sigma_to_Max + |lapply (MSC_in … CFf) #Hle + %{1} %{0} #n #_ @(transitive_le … (sigma_const …)) + >(commutative_times 1) (eq_minus_O … H) // + |lapply (le_S_S_to_le … (not_le_to_lt … H)) -H #H + @le_times // @(transitive_le … (Hle … )) + cut (b n = b n - a n + a n) [Hcut in ⊢ (?%?); @(le_Max (λi.sf 〈i+a n,n〉)) /2/ + ] + ] +qed. + +(* +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〉)). *) + + + \ No newline at end of file diff --git a/matita/matita/lib/reverse_complexity/complexity.ma b/matita/matita/lib/reverse_complexity/complexity.ma new file mode 100644 index 000000000..ae7baceca --- /dev/null +++ b/matita/matita/lib/reverse_complexity/complexity.ma @@ -0,0 +1,290 @@ +include "reverse_complexity/basics.ma". +include "reverse_complexity/big_O.ma". + +(********************************* complexity *********************************) + +(* We assume operations have a minimal structural complexity MSC. +For instance, for time complexity, MSC is equal to the size of input. +For space complexity, MSC is typically 0, since we only measure the space +required in addition to dimension of the input. *) + +axiom MSC : nat → nat. +axiom MSC_sublinear : ∀n. MSC (S n) ≤ S (MSC n). +(* axiom MSC_S: O MSC (λx.MSC (S x)) . *) +axiom MSC_le: ∀n. MSC n ≤ n. + +axiom monotonic_MSC: monotonic ? le MSC. +axiom MSC_pair_eq: ∀a,b. MSC 〈a,b〉 = MSC a + MSC b. + + +lemma MSC_pair: ∀a,b. MSC 〈a,b〉 ≤ MSC a + MSC b. // qed. + +(* C s i means i is running in O(s) *) + +definition C ≝ λs,i.∃c.∃a.∀x.a ≤ x → ∃y. + U i x (c*(s x)) = Some ? y. + +(* C f s means f ∈ O(s) where MSC ∈O(s) *) +definition CF ≝ λs,f.∃i.code_for (total f) i ∧ C s i. + +axiom MSC_in: ∀f,h. CF h f → ∀x. MSC x ≤ h x. +axiom MSC_out: ∀f,h. CF h f → ∀x. MSC (f x) ≤ h x. + + +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} whd in match (total ??); Hext @H | //] +qed. + +lemma monotonic_CF: ∀s1,s2,f.(∀x. s1 x ≤ s2 x) → CF s1 f → CF s2 f. +#s1 #s2 #f #H * #i * #Hcode #Hs1 +%{i} % [//] cases Hs1 #c * #a -Hs1 #Hs1 %{c} %{a} #n #lean +cases(Hs1 n lean) #y #Hy %{y} @(monotonic_U …Hy) @le_times // +qed. + +lemma O_to_CF: ∀s1,s2,f.O s2 s1 → CF s1 f → CF s2 f. +#s1 #s2 #f #H * #i * #Hcode #Hs1 +%{i} % [//] cases Hs1 #c * #a -Hs1 #Hs1 +cases H #c1 * #a1 #Ha1 %{(c*c1)} %{(a+a1)} #n #lean +cases(Hs1 n ?) [2:@(transitive_le … lean) //] #y #Hy %{y} @(monotonic_U …Hy) +>associative_times @le_times // @Ha1 @(transitive_le … lean) // +qed. + +lemma timesc_CF: ∀s,f,c.CF (λx.c*s x) f → CF s f. +#s #f #c @O_to_CF @O_times_c +qed. + +(********************************* composition ********************************) +axiom CF_comp: ∀f,g,sf,sg,sh. CF sg g → CF sf f → + O sh (λx. sg x + sf (g x)) → CF sh (f ∘ g). + +lemma CF_comp_ext: ∀f,g,h,sh,sf,sg. CF sg g → CF sf f → + (∀x.f(g x) = h x) → O sh (λx. sg x + sf (g x)) → CF sh h. +#f #g #h #sh #sf #sg #Hg #Hf #Heq #H @(ext_CF (f ∘ g)) + [#n normalize @Heq | @(CF_comp … H) //] +qed. + +(************************************* smn ************************************) +axiom smn: ∀f,s. CF s f → ∀x. CF (λy.s 〈x,y〉) (λy.f 〈x,y〉). + +(****************************** constructibility ******************************) + +definition constructible ≝ λs. CF s s. + +lemma constr_comp : ∀s1,s2. constructible s1 → constructible s2 → + (∀x. x ≤ s2 x) → constructible (s2 ∘ s1). +#s1 #s2 #Hs1 #Hs2 #Hle @(CF_comp … Hs1 Hs2) @O_plus @le_to_O #x [@Hle | //] +qed. + +lemma ext_constr: ∀s1,s2. (∀x.s1 x = s2 x) → + constructible s1 → constructible s2. +#s1 #s2 #Hext #Hs1 @(ext_CF … Hext) @(monotonic_CF … Hs1) #x >Hext // +qed. + + +(***************************** primitive recursion ****************************) + +(* no arguments *) + +let rec prim_rec0 (k:nat) (h:nat →nat) n on n ≝ + match n with + [ O ⇒ k + | S a ⇒ h 〈a, prim_rec0 k h a〉]. + +lemma prim_rec0_S: ∀k,h,n. + prim_rec0 k h (S n) = h 〈n, prim_rec0 k h n〉. +// qed. + +let rec prim_rec0_compl (k,sk:nat) (h,sh:nat →nat) n on n ≝ + match n with + [ O ⇒ sk + | S a ⇒ prim_rec0_compl k sk h sh a + sh (prim_rec0 k h a)]. + +axiom CF_prim_rec0_gen: ∀k,h,sk,sh,sh1,sf. CF sh h → + O sh1 (λx.snd x + sh 〈fst x,prim_rec0 k h (fst x)〉) → + O sf (prim_rec0 sk sh1) → + CF sf (prim_rec0 k h). + +lemma CF_prim_rec0: ∀k,h,sk,sh,sf. CF sh h → + O sf (prim_rec0 sk (λx. snd x + sh 〈fst x,prim_rec0 k h (fst x)〉)) + → CF sf (prim_rec0 k h). +#k #h #sk #sh #sf #Hh #HO @(CF_prim_rec0_gen … Hh … HO) @O_refl +qed. + +(* with arguments. m is a vector of arguments *) +let rec prim_rec (k,h:nat →nat) n m on n ≝ + match n with + [ O ⇒ k m + | S a ⇒ h 〈a,〈prim_rec k h a m, m〉〉]. + +lemma prim_rec_S: ∀k,h,n,m. + prim_rec k h (S n) m = h 〈n,〈prim_rec k h n m, m〉〉. +// qed. + +lemma prim_rec_le: ∀k1,k2,h1,h2. (∀x.k1 x ≤ k2 x) → +(∀x,y.x ≤y → h1 x ≤ h2 y) → + ∀x,m. prim_rec k1 h1 x m ≤ prim_rec k2 h2 x m. +#k1 #k2 #h1 #h2 #lek #leh #x #m elim x // +#n #Hind normalize @leh @le_pair // @le_pair // +qed. + +definition unary_pr ≝ λk,h,x. prim_rec k h (fst x) (snd x). + +lemma prim_rec_unary: ∀k,h,a,b. + prim_rec k h a b = unary_pr k h 〈a,b〉. +#k #h #a #b normalize >fst_pair >snd_pair // +qed. + +let rec prim_rec_compl (k,h,sk,sh:nat →nat) n m on n ≝ + match n with + [ O ⇒ sk m + | S a ⇒ prim_rec_compl k h sk sh a m + sh 〈a,〈prim_rec k h a m,m〉〉]. + +axiom CF_prim_rec_gen: ∀k,h,sk,sh,sh1. CF sk k → CF sh h → + O sh1 (λx. fst (snd x) + + sh 〈fst x,〈unary_pr k h 〈fst x,snd (snd x)〉,snd (snd x)〉〉) → + CF (unary_pr sk sh1) (unary_pr k h). + +lemma CF_prim_rec: ∀k,h,sk,sh,sf. CF sk k → CF sh h → + O sf (unary_pr sk + (λx. fst (snd x) + + sh 〈fst x,〈unary_pr k h 〈fst x,snd (snd x)〉,snd (snd x)〉〉)) + → CF sf (unary_pr k h). +#k #h #sk #sh #sf #Hk #Hh #Osf @(O_to_CF … Osf) @(CF_prim_rec_gen … Hk Hh) +@O_refl +qed. + +lemma constr_prim_rec: ∀s1,s2. constructible s1 → constructible s2 → + (∀n,r,m. 2 * r ≤ s2 〈n,〈r,m〉〉) → constructible (unary_pr s1 s2). +#s1 #s2 #Hs1 #Hs2 #Hincr @(CF_prim_rec … Hs1 Hs2) whd %{2} %{0} +#x #_ lapply (surj_pair x) * #a * #b #eqx >eqx whd in match (unary_pr ???); +>fst_pair >snd_pair +whd in match (unary_pr ???); >fst_pair >snd_pair elim a + [normalize // + |#n #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair + >prim_rec_S @transitive_le [| @(monotonic_le_plus_l … Hind)] + @transitive_le [| @(monotonic_le_plus_l … (Hincr n ? b))] + whd in match (unary_pr ???); >fst_pair >snd_pair // + ] +qed. + +(**************************** primitive operations*****************************) + +definition id ≝ λx:nat.x. + +axiom CF_id: CF MSC id. +axiom CF_const: ∀k.CF MSC (λx.k). +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_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. + +(**************************** arithmetic operations ***************************) + +axiom CF_compS: ∀h,f. CF h f → CF h (S ∘ f). +axiom CF_le: ∀h,f,g. CF h f → CF h g → CF h (λx. leb (f x) (g x)). +axiom CF_eqb: ∀h,f,g. CF h f → CF h g → CF h (λx.eqb (f x) (g x)). +axiom CF_max1: ∀h,f,g. CF h f → CF h g → CF h (λx. max (f x) (g x)). +axiom CF_plus: ∀h,f,g. CF h f → CF h g → CF h (λx. f x + g x). +axiom CF_minus: ∀h,f,g. CF h f → CF h g → CF h (λx. f x - g x). + +(******************************** if then else ********************************) + +lemma if_prim_rec: ∀b:nat → bool. ∀f,g:nat → nat.∀x:nat. + if b x then f x else g x = prim_rec g (f ∘ snd ∘ snd) (b x) x. +#b #f #g #x cases (b x) normalize // +qed. + +lemma CF_if: ∀b:nat → bool. ∀f,g,s. CF s b → CF s f → CF s g → + CF s (λx. if b x then f x else g x). +#b #f #g #s #CFb #CFf #CFg @(ext_CF (λx.unary_pr g (f ∘ snd ∘ snd) 〈b x,x〉)) + [#n @sym_eq normalize >fst_pair >snd_pair @if_prim_rec + |@(CF_comp ??? s) + [|@CF_comp_pair // @(O_to_CF … CF_id) @le_to_O @MSC_in // + |@(CF_prim_rec_gen ??? (λx.MSC x + s(snd (snd x))) … CFg) [3:@O_refl|] + @(CF_comp … (λx.MSC x + s(snd x)) … CF_snd) + [@(CF_comp … s … CF_snd CFf) @O_refl + |@O_plus + [@O_plus_l @O_refl + |@O_plus + [@O_plus_l @le_to_O #x @monotonic_MSC // + |@O_plus_r @O_refl + ] + ] + ] + |%{6} %{0} #n #_ normalize in ⊢ (??%); snd_pair >fst_pair normalize + [@le_plus [//] >fst_pair >fst_pair >snd_pair >fst_pair + @le_plus [//] >snd_pair >snd_pair >snd_pair >snd_pair + associative_plus @le_plus + [@(transitive_le ???? (MSC_in … CFf n)) @monotonic_MSC // + |@(transitive_le … (MSC_pair …)) @le_plus [|@(MSC_in … CFf)] + normalize @MSC_out @CFg + ] + |@le_plus // + ] + ] + ] +qed. + +(********************************* simulation *********************************) + +definition pU_unary ≝ λp. pU (fst p) (fst (snd p)) (snd (snd p)). + +axiom sU : nat → nat. +axiom CF_U : CF sU pU_unary. + +axiom monotonic_sU: ∀i1,i2,x1,x2,s1,s2. i1 ≤ i2 → x1 ≤ x2 → s1 ≤ s2 → + sU 〈i1,〈x1,s1〉〉 ≤ sU 〈i2,〈x2,s2〉〉. + +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) +#b1 * #c1 #eqy >eqy -eqy +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_pos: ∀x. 0 < sU x. + +axiom sU_le: ∀i,x,s. s ≤ sU 〈i,〈x,s〉〉. + +lemma sU_le_i: ∀i,x,s. MSC i ≤ sU 〈i,〈x,s〉〉. +#i #x #s @(transitive_le ???? (MSC_in … CF_U 〈i,〈x,s〉〉)) @monotonic_MSC // +qed. + +lemma sU_le_x: ∀i,x,s. MSC x ≤ sU 〈i,〈x,s〉〉. +#i #x #s @(transitive_le ???? (MSC_in … CF_U 〈i,〈x,s〉〉)) @monotonic_MSC +@(transitive_le … 〈x,s〉) // +qed. + + + + +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. + diff --git a/matita/matita/lib/reverse_complexity/gap.ma b/matita/matita/lib/reverse_complexity/gap.ma new file mode 100644 index 000000000..2c8b6e938 --- /dev/null +++ b/matita/matita/lib/reverse_complexity/gap.ma @@ -0,0 +1,251 @@ + +include "arithmetics/minimization.ma". +include "arithmetics/bigops.ma". +include "arithmetics/pidgeon_hole.ma". +include "arithmetics/iteration.ma". + +(************************** notation for miminimization ***********************) + +(* an alternative defintion of minimization +definition Min ≝ λa,f. + \big[min,a]_{i < a | f i} i. *) + +notation "μ_{ ident i < n } p" + with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ≤ n } p" + with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b] } p" + with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. + +lemma f_min_true: ∀f,a,b. + (∃i. a ≤ i ∧ i ≤ b ∧ f i = true) → f (μ_{i ∈[a,b]} (f i)) = true. +#f #a #b * #i * * #Hil #Hir #Hfi @(f_min_true … (λx. f x)) Hcut in ⊢ (??%); @lt_min %{i} % // % [@Hil |Hm #HS destruct (HS) // + |#ltmn lapply (monotonic_U … n … Hm) [@lt_to_le @not_le_to_lt //] + >Hn #HS destruct (HS) // + ] +qed. + +definition terminate ≝ λi,x,r. ∃y. U i x r = Some ? y. + +notation "〈i,x〉 ↓ r" with precedence 60 for @{terminate $i $x $r}. + +lemma terminate_dec: ∀i,x,n. 〈i,x〉 ↓ n ∨ ¬ 〈i,x〉 ↓ n. +#i #x #n normalize cases (U i x n) + [%2 % * #y #H destruct|#y %1 %{y} //] +qed. + +definition termb ≝ λi,x,t. + match U i x t with [None ⇒ false |Some y ⇒ true]. + +lemma termb_true_to_term: ∀i,x,t. termb i x t = true → 〈i,x〉 ↓ t. +#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] +qed. + +lemma term_to_termb_true: ∀i,x,t. 〈i,x〉 ↓ t → termb i x t = true. +#i #x #t * #y #H normalize >H // +qed. + +lemma decidable_test : ∀n,x,r,r1. + (∀i. i < n → 〈i,x〉 ↓ r ∨ ¬ 〈i,x〉 ↓ r1) ∨ + (∃i. i < n ∧ (¬ 〈i,x〉 ↓ r ∧ 〈i,x〉 ↓ r1)). +#n #x #r1 #r2 + cut (∀i0.decidable ((〈i0,x〉↓r1) ∨ ¬ 〈i0,x〉 ↓ r2)) + [#j @decidable_or [@terminate_dec |@decidable_not @terminate_dec ]] #Hdec + cases(decidable_forall ? Hdec n) + [#H %1 @H + |#H %2 cases (not_forall_to_exists … Hdec H) #j * #leji #Hj + %{j} % // % + [@(not_to_not … Hj) #H %1 @H + |cases (terminate_dec j x r2) // #H @False_ind cases Hj -Hj #Hj + @Hj %2 @H + ] +qed. + +(**************************** the gap theorem *********************************) +definition gapP ≝ λn,x,g,r. ∀i. i < n → 〈i,x〉 ↓ r ∨ ¬ 〈i,x〉 ↓ g r. + +lemma gapP_def : ∀n,x,g,r. + gapP n x g r = ∀i. i < n → 〈i,x〉 ↓ r ∨ ¬ 〈i,x〉 ↓ g r. +// qed. + +lemma upper_bound_aux: ∀g,b,n,x. (∀x. x ≤ g x) → ∀k. + (∃j.j < k ∧ + (∀i. i < n → 〈i,x〉 ↓ g^j b ∨ ¬ 〈i,x〉 ↓ g^(S j) b)) ∨ + ∃l. |l| = k ∧ unique ? l ∧ ∀i. i ∈ l → i < n ∧ 〈i,x〉 ↓ g^k b . +#g#b #n #x #Hg #k elim k + [%2 %{([])} normalize % [% //|#x @False_ind] + |#k0 * + [* #j * #lej #H %1 %{j} % [@le_S // | @H ] + |* #l * * #Hlen #Hunique #Hterm + cases (decidable_test n x (g^k0 b) (g^(S k0) b)) + [#Hcase %1 %{k0} % [@le_n | @Hcase] + |* #j * #ltjn * #H1 #H2 %2 + %{(j::l)} % + [ % [normalize @eq_f @Hlen] whd % // % #H3 + @(absurd ?? H1) @(proj2 … (Hterm …)) @H3 + |#x * + [#eqxj >eqxj % // + |#Hmemx cases(Hterm … Hmemx) #lexn * #y #HU + % [@lexn] %{y} @(monotonic_U ?????? HU) @Hg + ] + ] + ] + ] + ] +qed. + +lemma upper_bound: ∀g,b,n,x. (∀x. x ≤ g x) → ∃r. + (* b ≤ r ∧ r ≤ g^n b ∧ ∀i. i < n → 〈i,x〉 ↓ r ∨ ¬ 〈i,x〉 ↓ g r. *) + b ≤ r ∧ r ≤ g^n b ∧ gapP n x g r. +#g #b #n #x #Hg +cases (upper_bound_aux g b n x Hg n) + [* #j * #Hj #H %{(g^j b)} % [2: @H] % [@le_iter //] + @monotonic_iter2 // @lt_to_le // + |* #l * * #Hlen #Hunique #Hterm %{(g^n b)} % + [% [@le_iter // |@le_n]] + #i #lein %1 @(proj2 … (Hterm ??)) + @(eq_length_to_mem_all … Hlen Hunique … lein) + #x #memx @(proj1 … (Hterm ??)) // + ] +qed. + +definition gapb ≝ λn,x,g,r. + \big[andb,true]_{i < n} ((termb i x r) ∨ ¬(termb i x (g r))). + +lemma gapb_def : ∀n,x,g,r. gapb n x g r = + \big[andb,true]_{i < n} ((termb i x r) ∨ ¬(termb i x (g r))). +// qed. + +lemma gapb_true_to_gapP : ∀n,x,g,r. + gapb n x g r = true → ∀i. i < n → 〈i,x〉 ↓ r ∨ ¬(〈i,x〉 ↓ (g r)). +#n #x #g #r elim n + [>gapb_def >bigop_Strue // + #H #i #lti0 @False_ind @(absurd … lti0) @le_to_not_lt // + |#m #Hind >gapb_def >bigop_Strue // + #H #i #leSm cases (le_to_or_lt_eq … leSm) + [#lem @Hind [@(andb_true_r … H)|@le_S_S_to_le @lem] + |#eqi >(injective_S … eqi) lapply (andb_true_l … H) -H #H cases (orb_true_l … H) -H + [#H %1 @termb_true_to_term // + |#H %2 % #H1 >(term_to_termb_true … H1) in H; normalize #H destruct + ] + ] + ] +qed. + +lemma gapP_to_gapb_true : ∀n,x,g,r. + (∀i. i < n → 〈i,x〉 ↓ r ∨ ¬(〈i,x〉 ↓ (g r))) → gapb n x g r = true. +#n #x #g #r elim n // +#m #Hind #H >gapb_def >bigop_Strue // @true_to_andb_true + [cases (H m (le_n …)) + [#H2 @orb_true_r1 @term_to_termb_true // + |#H2 @orb_true_r2 @sym_eq @noteq_to_eqnot @sym_not_eq + @(not_to_not … H2) @termb_true_to_term + ] + |@Hind #i0 #lei0 @H @le_S // + ] +qed. + + +(* the gap function *) +let rec gap g n on n ≝ + match n with + [ O ⇒ 1 + | S m ⇒ let b ≝ gap g m in μ_{i ∈ [b,g^n b]} (gapb n n g i) + ]. + +lemma gapS: ∀g,m. + gap g (S m) = + (let b ≝ gap g m in + μ_{i ∈ [b,g^(S m) b]} (gapb (S m) (S m) g i)). +// qed. + +lemma upper_bound_gapb: ∀g,m. (∀x. x ≤ g x) → + ∃r:ℕ.gap g m ≤ r ∧ r ≤ g^(S m) (gap g m) ∧ gapb (S m) (S m) g r = true. +#g #m #leg +lapply (upper_bound g (gap g m) (S m) (S m) leg) * #r * * +#H1 #H2 #H3 %{r} % + [% // |@gapP_to_gapb_true @H3] +qed. + +lemma gapS_true: ∀g,m. (∀x. x ≤g x) → gapb (S m) (S m) g (gap g (S m)) = true. +#g #m #leg @(f_min_true (gapb (S m) (S m) g)) @upper_bound_gapb // +qed. + +theorem gap_theorem: ∀g,i.(∀x. x ≤ g x)→∃k.∀n.k < n → + 〈i,n〉 ↓ (gap g n) ∨ ¬ 〈i,n〉 ↓ (g (gap g n)). +#g #i #leg %{i} * + [#lti0 @False_ind @(absurd ?? (not_le_Sn_O i) ) // + |#m #leim lapply (gapS_true g m leg) #H + @(gapb_true_to_gapP … H) // + ] +qed. + +(* an upper bound *) + +let rec sigma n ≝ + match n with + [ O ⇒ 0 | S m ⇒ n + sigma m ]. + +lemma gap_bound: ∀g. (∀x. x ≤ g x) → (monotonic ? le g) → + ∀n.gap g n ≤ g^(sigma n) 1. +#g #leg #gmono #n elim n + [normalize // + |#m #Hind >gapS @(transitive_le ? (g^(S m) (gap g m))) + [@min_up @upper_bound_gapb // + |@(transitive_le ? (g^(S m) (g^(sigma m) 1))) + [@monotonic_iter // |>iter_iter >commutative_plus @le_n + ] + ] +qed. + +lemma gap_bound2: ∀g. (∀x. x ≤ g x) → (monotonic ? le g) → + ∀n.gap g n ≤ g^(n*n) 1. +#g #leg #gmono #n elim n + [normalize // + |#m #Hind >gapS @(transitive_le ? (g^(S m) (gap g m))) + [@min_up @upper_bound_gapb // + |@(transitive_le ? (g^(S m) (g^(m*m) 1))) + [@monotonic_iter // + |>iter_iter @monotonic_iter2 [@leg | normalize 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. + +alias id "max" = "cic:/matita/arithmetics/nat/max#def:2". +alias id "mk_Aop" = "cic:/matita/arithmetics/bigops/Aop#con:0:1:2". +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_plus_r: ∀f,s1,s2. O s2 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_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 (times_n_1 (s n0)) in ⊢ (?%?); >commutative_times @H // +qed. + +lemma o_trans: ∀s1,s2,s3. o s2 s1 → o s3 s2 → o s3 s1. +#s1 #s2 #s3 #H1 #H2 #c cases (H1 c) #n1 -H1 #H1 cases (H2 1) #n2 -H2 #H2 +%{(max n1 n2)} #n #Hmax +@(transitive_lt … (H1 ??)) [@(le_maxl … Hmax)] +>(times_n_1 (s2 n)) in ⊢ (?%?); >commutative_times @H2 @(le_maxr … Hmax) +qed. + + +(*********************************** pairing **********************************) + +axiom pair: nat →nat →nat. +axiom fst : nat → nat. +axiom snd : nat → nat. +axiom fst_pair: ∀a,b. fst (pair a b) = a. +axiom snd_pair: ∀a,b. snd (pair a b) = b. + +interpretation "abstract pair" 'pair f g = (pair f g). + +(************************ basic complexity notions ****************************) + +axiom U: nat → nat → nat → option nat. + +axiom monotonic_U: ∀i,x,n,m,y.n ≤m → + U i x n = Some ? y → U i x m = Some ? y. + +lemma unique_U: ∀i,x,n,m,yn,ym. + U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. +#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. +notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $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 monotonic_size: monotonic ? le size. + +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. + +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. + +(* C s i means that the complexity of i is O(s) *) + +definition C ≝ λs,i.∃c.∃a.∀x.a ≤ |x| → {i ⊙ x} ↓ (c*(s(|x|))). + +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 hierachy theorem (left) **************************) + +theorem hierarchy_theorem_left: ∀s1,s2:nat→nat. + O(s1) ⊆ O(s2) → CF s1 ⊆ CF s2. +#s1 #s2 #HO #f * #i * #Hcode * #c * #a #Hs1_i %{i} % // +cases (sub_O_to_O … HO) -HO #c1 * #b #Hs1s2 +%{(c*c1)} %{(max a b)} #x #lemax +cases (Hs1_i x ?) [2: @(le_maxl …lemax)] +#y #Hy %{y} @(monotonic_U … Hy) >associative_times +@le_times // @Hs1s2 @(le_maxr … lemax) +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. + +let rec f_img (f:nat →nat) n on n ≝ + match n with + [O ⇒ [ ] + |S m ⇒ f m::f_img f m + ]. + +(* a few lemma to prove injective_to_exists. This is a general result; a nice +example of the pidgeon hole pricniple *) + +lemma f_img_to_exists: + ∀f.∀n,a. a ∈ f_img f n → ∃b. b < n ∧ f b = a. +#f #n #a elim n normalize [@False_ind] +#m #Hind * + [#Ha %{m} /2/ |#H cases(Hind H) #b * #Hb #Ha %{b} % // @le_S //] +qed. + +lemma length_f_img: ∀f,n. |f_img f n| = n. +#f #n elim n // normalize // +qed. + +lemma unique_f_img: ∀f,n. injective … f → unique ? (f_img f n). +#f #n #Hinj elim n normalize // +#m #Hind % // % #H lapply(f_img_to_exists …H) * #b * #ltbm +#eqbm @(absurd … ltbm) @le_to_not_lt >(Hinj… eqbm) // +qed. + +lemma injective_to_exists: ∀f. injective nat nat f → + ∀n.(∀i.i < n → f i < n) → ∀a. a < n → ∃b. bHxn in Hx; /2 by absurd/ + |#H lapply(not_forall_to_exists … Hdec H) + * #b * #H1 #H2 %{b} @not_lt_to_le @H2 + ] +qed. + +lemma pad : ∀n,a. ∃b. n ≤ |〈a,b〉|. +#n #a cases (weak_pad1 (of_size n) a) #b #Hb +%{b} <(size_of_size n) @monotonic_size // +qed. + +lemma o_to_ex: ∀s1,s2. o s1 s2 → ∀i. C s2 i → + ∃b.{i ⊙ 〈i,b〉} ↓ s1 (|〈i,b〉|). +#s1 #s2 #H #i * #c * #x0 #H1 +cases (H c) #n0 #H2 cases (pad (max x0 n0) i) #b #Hmax +%{b} cases (H1 〈i,b〉 ?) + [#z #H3 %{z} @(monotonic_U … H3) @lt_to_le @H2 + @(le_maxr … Hmax) + |@(le_maxl … Hmax) + ] +qed. + +lemma diag1_not_s1: ∀s1,s2. o s1 s2 → ¬ CF s2 (diag_cf s1). +#s1 #s2 #H1 % * #i * #Hcode_i #Hs2_i +cases (o_to_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 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 sU2: nat → nat → nat. +axiom sU: nat → nat → nat → nat. + +axiom CFU_new: ∀h,g,f,s. + CF s (to_Some h) → CF s (to_Some g) → CF s (to_Some f) → + O s (λx. sU (size_f h x) (size_f g x) (size_f f x)) → + CF s (λx.U (h x) (g x) (|f x|)). + +lemma 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|)). +#h #g #f #s1 #s2 #s3 #Hh #Hg #Hf @CFU_new + [@(monotonic_CF … Hh) @O_plus_l @O_plus_l @O_plus_l // + |@(monotonic_CF … Hg) @O_plus_l @O_plus_l @O_plus_r // + |@(monotonic_CF … Hf) @O_plus_l @O_plus_r // + |@O_plus_r // + ] +qed. + +axiom monotonic_sU: ∀a1,a2,b1,b2,c1,c2. a1 ≤ a2 → b1 ≤ b2 → c1 ≤c2 → + sU a1 b1 c1 ≤ sU a2 b2 c2. + +axiom superlinear_sU: ∀i,x,r. r ≤ sU i x r. + +(* not used *) +definition sU_space ≝ λi,x,r.i+x+r. +definition sU_time ≝ λi,x,r.i+x+(i^2)*r*(log 2 r). + +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_new: ∀b,f,g,s. CF s b → CF s f → CF s g → CF s (IF b f g). + +lemma 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). +#b #f #g #sb #sf #sg #Hb #Hf #Hg @IF_CF_new + [@(monotonic_CF … Hb) @O_plus_l @O_plus_l // + |@(monotonic_CF … Hf) @O_plus_l @O_plus_r // + |@(monotonic_CF … Hg) @O_plus_r // + ] +qed. + +lemma diag_cf_def : ∀s.∀i. + diag_cf s i = + IF (λi.U (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 sufficiently_large ≝ λs. CF s (λn. Some ? n) ∧ ∀c. CF s (λn. Some ? c). + +definition constructible ≝ λs. CF s (λx.Some ? (of_size (s (|x|)))). + +lemma diag_s: ∀s. sufficiently_large s → constructible s → + CF (λx.sU x x (s x)) (diag_cf s). +#s * #Hs_id #Hs_c #Hs_constr +cut (O (λx:ℕ.sU x x (s x)) s) [%{1} %{0} #n //] +#O_sU_s @ext_CF [2: #n @sym_eq @diag_cf_def | skip] +@IF_CF_new [2,3:@(monotonic_CF … (Hs_c ?)) // ] +@CFU_new + [@CF_fst @(monotonic_CF … Hs_id) // + |@(monotonic_CF … Hs_id) // + |@(monotonic_CF … Hs_constr) // + |%{1} %{0} #n #_ >commutative_times size_f_size >size_of_size // + ] +qed. \ No newline at end of file diff --git a/matita/matita/lib/reverse_complexity/sigma_diseq.ma b/matita/matita/lib/reverse_complexity/sigma_diseq.ma new file mode 100644 index 000000000..4356cd24e --- /dev/null +++ b/matita/matita/lib/reverse_complexity/sigma_diseq.ma @@ -0,0 +1,153 @@ +include "arithmetics/sigma_pi.ma". + +(************************* notation for minimization **************************) +notation "μ_{ ident i < n } p" + with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ≤ n } p" + with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b[ } p" + with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b] } p" + with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. + +(************************************ 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 le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true → + f a ≤ Max_{i ∈ [m,n[ | p i}(f i). +#f #p #n #m #a #lema #ltan #pa +>(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m)) + [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. + + +(************************* a couple of technical lemmas ***********************) +lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0. +#a elim a // #n #Hind * + [#H @False_ind /2 by absurd/ | #b normalize #H @Hind @le_S_S_to_le /2/] +qed. + +lemma sigma_const: ∀c,a,b. + ∑_{i ∈ [a,b[ }c ≤ (b-a)*c. +#c #a #b elim (b-a) // #n #Hind normalize @le_plus // +qed. + +lemma sigma_to_Max: ∀h,a,b. + ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*Max_{i ∈ [a,b[ }(h i). +#h #a #b elim (b-a) + [// + |#n #Hind >bigop_Strue [2://] whd in ⊢ (??%); + @le_plus + [>bigop_Strue // @(le_maxl … (le_n …)) + |@(transitive_le … Hind) @le_times // >bigop_Strue // + @(le_maxr … (le_n …)) + ] + ] +qed. + +lemma sigma_bound1: ∀h,a,b. monotonic nat le h → + ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h b. +#h #a #b #H +@(transitive_le … (sigma_to_Max …)) @le_times // +@Max_le #i #lti #_ @H @lt_to_le @lt_minus_to_plus_r // +qed. + +lemma sigma_bound_decr1: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) → + ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a. +#h #a #b #H +@(transitive_le … (sigma_to_Max …)) @le_times // +@Max_le #i #lti #_ @H // @lt_minus_to_plus_r // +qed. + +lemma sigma_bound: ∀h,a,b. monotonic nat le h → + ∑_{i ∈ [a,S b[ }(h i) ≤ (S b-a)*h b. +#h #a #b #H cases (decidable_le a b) + [#leab cut (b = pred (S b - a + a)) + [Hb in match (h b); + generalize in match (S b -a); + #n elim n + [// + |#m #Hind >bigop_Strue [2://] @le_plus + [@H @le_n |@(transitive_le … Hind) @le_times [//] @H //] + ] + |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba + cut (S b -a = 0) [@minus_to_0 //] #Hcut >Hcut // + ] +qed. + +lemma sigma_bound_decr: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) → + ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a. +#h #a #b #H cases (decidable_le a b) + [#leab cut ((b -a) +a ≤ b) [/2 by le_minus_to_plus_r/] generalize in match (b -a); + #n elim n + [// + |#m #Hind >bigop_Strue [2://] #Hm + cut (m+a ≤ b) [@(transitive_le … Hm) //] #Hm1 + @le_plus [@H // |@(transitive_le … (Hind Hm1)) //] + ] + |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba + cut (b -a = 0) [@minus_to_0 @lt_to_le @ltba] #Hcut >Hcut // + ] +qed. + \ No newline at end of file diff --git a/matita/matita/lib/reverse_complexity/speed_clean.ma b/matita/matita/lib/reverse_complexity/speed_clean.ma new file mode 100644 index 000000000..bfd3d34b1 --- /dev/null +++ b/matita/matita/lib/reverse_complexity/speed_clean.ma @@ -0,0 +1,1068 @@ +include "basics/types.ma". +include "arithmetics/minimization.ma". +include "arithmetics/bigops.ma". +include "arithmetics/sigma_pi.ma". +include "arithmetics/bounded_quantifiers.ma". +include "reverse_complexity/big_O.ma". +include "basics/core_notation/napart_2.ma". + +(************************* notation for minimization *****************************) +notation "μ_{ ident i < n } p" + with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ≤ n } p" + with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b[ } p" + with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b] } p" + with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. + +(************************************ 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 le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true → + f a ≤ Max_{i ∈ [m,n[ | p i}(f i). +#f #p #n #m #a #lema #ltan #pa +>(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m)) + [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. + +(********************************** pairing ***********************************) +axiom pair: nat → nat → nat. +axiom fst : nat → nat. +axiom snd : nat → nat. + +interpretation "abstract pair" 'pair f g = (pair f g). + +axiom fst_pair: ∀a,b. fst 〈a,b〉 = a. +axiom snd_pair: ∀a,b. snd 〈a,b〉 = b. +axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉. + +axiom le_fst : ∀p. fst p ≤ p. +axiom le_snd : ∀p. snd p ≤ p. +axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉. + +(************************************* U **************************************) +axiom U: nat → nat →nat → option nat. + +axiom monotonic_U: ∀i,x,n,m,y.n ≤m → + U i x n = Some ? y → U i x m = Some ? y. + +lemma unique_U: ∀i,x,n,m,yn,ym. + U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. +#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. + +notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. + +lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n. +#i #x #n normalize cases (U i x n) + [%2 % * #y #H destruct|#y %1 %{y} //] +qed. + +lemma monotonic_terminate: ∀i,x,n,m. + n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m. +#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) // +qed. + +definition termb ≝ λi,x,t. + match U i x t with [None ⇒ false |Some y ⇒ true]. + +lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t. +#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] +qed. + +lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true. +#i #x #t * #y #H normalize >H // +qed. + +definition out ≝ λi,x,r. + match U i x r with [ None ⇒ 0 | Some z ⇒ z]. + +definition bool_to_nat: bool → nat ≝ + λb. match b with [true ⇒ 1 | false ⇒ 0]. + +coercion bool_to_nat. + +definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉. + +lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y. +#i #x #r #y % normalize + [cases (U i x r) normalize + [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H] + #H1 destruct + |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1] + #H1 // + ] + |#H >H //] +qed. + +lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?. +#i #x #r % normalize + [cases (U i x r) normalize // + #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1] + #H1 destruct + |#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)). + +lemma min_input_def : ∀h,i,x. + min_input h i x = min (x -i) (S i) (λy.termb i y (h (S i) y)). +// qed. + +lemma min_input_i: ∀h,i,x. x ≤ i → min_input h i x = S i. +#h #i #x #lexi >min_input_def +cut (x - i = 0) [@sym_eq /2 by eq_minus_O/] #Hcut // +qed. + +lemma min_input_to_terminate: ∀h,i,x. + min_input h i x = x → {i ⊙ x} ↓ (h (S i) x). +#h #i #x #Hminx +cases (decidable_le (S i) x) #Hix + [cases (true_or_false (termb i x (h (S i) x))) #Hcase + [@termb_true_to_term // + |min_input_def in Hminx; #Hminx >Hminx in ⊢ (%→?); + min_input_i in Hminx; + [#eqix >eqix in Hix; * /2/ | @le_S_S_to_le @not_le_to_lt //] + ] +qed. + +lemma min_input_to_lt: ∀h,i,x. + min_input h i x = x → i < x. +#h #i #x #Hminx cases (decidable_le (S i) x) // +#ltxi @False_ind >min_input_i in Hminx; + [#eqix >eqix in ltxi; * /2/ | @le_S_S_to_le @not_le_to_lt //] +qed. + +lemma le_to_min_input: ∀h,i,x,x1. x ≤ x1 → + min_input h i x = x → min_input h i x1 = x. +#h #i #x #x1 #lex #Hminx @(min_exists … (le_S_S … lex)) + [@(fmin_true … (sym_eq … Hminx)) // + |@(min_input_to_lt … Hminx) + |#j #H1 g_def cut (x-u = 0) [/2 by minus_le_minus_minus_comm/] +#eq0 >eq0 normalize // qed. + +lemma g_lt : ∀h,i,x. min_input h i x = x → + out i x (h (S i) x) < g h 0 x. +#h #i #x #H @le_S_S @(le_MaxI … i) /2 by min_input_to_lt/ +qed. + +(* +axiom ax1: ∀h,i. + (∃y.i < y ∧ (termb i y (h (S i) y)=true)) ∨ + ∀y. i < y → (termb i y (h (S i) y)=false). + +lemma eventually_0: ∀h,u.∃nu.∀x. nu < x → + max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) = 0. +#h #u elim u + [%{0} normalize // + |#u0 * #nu0 #Hind cases (ax1 h u0) + [* #x0 * #leu0x0 #Hx0 %{(max nu0 x0)} + #x #Hx >bigop_Sfalse + [>(minus_n_O u0) @Hind @(le_to_lt_to_lt … Hx) /2 by le_maxl/ + |@not_eq_to_eqb_false % #Hf @(absurd (x ≤ x0)) + [bigop_Sfalse + [>(minus_n_O u0) @Hind @(le_to_lt_to_lt … Hx) @le_maxr // + |@not_eq_to_eqb_false >min_input_def + >(min_not_exists (λy.(termb (u0+0) y (h (S (u0+0)) y)))) + [(bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA) + [>H // @(le_to_lt_to_lt …Hx) /2 by le_maxl/ + |@lt_to_le @(le_to_lt_to_lt …Hx) /2 by le_maxr/ + |// + ] +qed. *) + +lemma max_neq0 : ∀a,b. max a b ≠ 0 → a ≠ 0 ∨ b ≠ 0. +#a #b whd in match (max a b); cases (true_or_false (leb a b)) #Hcase >Hcase + [#H %2 @H | #H %1 @H] +qed. + +definition almost_equal ≝ λf,g:nat → nat. ¬ ∀nu.∃x. nu < x ∧ f x ≠ g x. +interpretation "almost equal" 'napart f g = (almost_equal f g). + +lemma eventually_cancelled: ∀h,u.¬∀nu.∃x. nu < x ∧ + max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) ≠ 0. +#h #u elim u + [normalize % #H cases (H u) #x * #_ * #H1 @H1 // + |#u0 @not_to_not #Hind #nu cases (Hind nu) #x * #ltx + cases (true_or_false (eqb (min_input h (u0+O) x) x)) #Hcase + [>bigop_Strue [2:@Hcase] #Hmax cases (max_neq0 … Hmax) -Hmax + [2: #H %{x} % // bigop_Sfalse + [#H %{x1} % [@transitive_lt //| (le_to_min_input … (eqb_true_to_eq … Hcase)) + [@lt_to_not_eq @ltx1 | @lt_to_le @ltx1] + ] + |>bigop_Sfalse [2:@Hcase] #H %{x} % // (bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA) + [>H // |@lt_to_le @(le_to_lt_to_lt …ltx) /2 by le_maxr/ |//] +qed. + +(******************************** Condition 2 *********************************) +definition total ≝ λf.λx:nat. Some nat (f x). + +lemma exists_to_exists_min: ∀h,i. (∃x. i < x ∧ {i ⊙ x} ↓ h (S i) x) → ∃y. min_input h i y = y. +#h #i * #x * #ltix #Hx %{(min_input h i x)} @min_spec_to_min @found // + [@(f_min_true (λy:ℕ.termb i y (h (S i) y))) %{x} % [% // | @term_to_termb_true //] + |#y #leiy #lty @(lt_min_to_false ????? lty) // + ] +qed. + +lemma condition_2: ∀h,i. code_for (total (g h 0)) i → ¬∃x. itermy >Hr +#H @(absurd ? H) @le_to_not_lt @le_n +qed. + + +(********************** complexity ***********************) + +(* We assume operations have a minimal structural complexity MSC. +For instance, for time complexity, MSC is equal to the size of input. +For space complexity, MSC is typically 0, since we only measure the +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) *) + +definition C ≝ λs,i.∃c.∃a.∀x.a ≤ x → ∃y. + U i x (c*(s x)) = Some ? 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 ??); (H m leam) normalize @eq_f @Hext +qed. *) + +lemma monotonic_CF: ∀s1,s2,f.(∀x. s1 x ≤ s2 x) → CF s1 f → CF s2 f. +#s1 #s2 #f #H * #HO * #i * #Hcode #Hs1 % + [cases HO #c * #a -HO #HO %{c} %{a} #n #lean @(transitive_le … (HO n lean)) + @le_times // + |%{i} % [//] cases Hs1 #c * #a -Hs1 #Hs1 %{c} %{a} #n #lean + cases(Hs1 n lean) #y #Hy %{y} @(monotonic_U …Hy) @le_times // + ] +qed. + +lemma O_to_CF: ∀s1,s2,f.O s2 s1 → CF s1 f → CF s2 f. +#s1 #s2 #f #H * #HO * #i * #Hcode #Hs1 % + [@(O_trans … H) // + |%{i} % [//] cases Hs1 #c * #a -Hs1 #Hs1 + cases H #c1 * #a1 #Ha1 %{(c*c1)} %{(a+a1)} #n #lean + cases(Hs1 n ?) [2:@(transitive_le … lean) //] #y #Hy %{y} @(monotonic_U …Hy) + >associative_times @le_times // @Ha1 @(transitive_le … lean) // + ] +qed. + +lemma timesc_CF: ∀s,f,c.CF (λx.c*s x) f → CF s f. +#s #f #c @O_to_CF @O_times_c +qed. + +(********************************* composition ********************************) +axiom CF_comp: ∀f,g,sf,sg,sh. CF sg g → CF sf f → + O sh (λx. sg x + sf (g x)) → CF sh (f ∘ g). + +lemma CF_comp_ext: ∀f,g,h,sh,sf,sg. CF sg g → CF sf f → + (∀x.f(g x) = h x) → O sh (λx. sg x + sf (g x)) → CF sh h. +#f #g #h #sh #sf #sg #Hg #Hf #Heq #H @(ext_CF (f ∘ g)) + [#n normalize @Heq | @(CF_comp … H) //] +qed. + +(* +lemma CF_comp1: ∀f,g,s. CF s (total g) → CF s (total f) → + CF s (total (f ∘ g)). +#f #g #s #Hg #Hf @(timesc_CF … 2) @(monotonic_CF … (CF_comp … Hg Hf)) +*) + +(* +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). + +lemma main_MSC: ∀h,f. CF h f → O h (λx.MSC (f x)). + +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 f → CF h (S ∘ f). +#h #f #Hf @(CF_comp … Hf CF_S) @O_plus // @main_MSC // +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 // +qed. + +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_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 ≝ + λf.λx:nat. match f x with [true ⇒ Some ? 0 |false ⇒ Some ? 1]. *) + +axiom CF_eqb: ∀h,f,g. + CF h f → CF h g → CF h (λx.eqb (f x) (g x)). + +(* +axiom eqb_compl2: ∀h,f,g. + CF2 h (total2 f) → CF2 h (total2 g) → + CF2 h (btotal2 (λx1,x2.eqb (f x1 x2) (g x1 x2))). + +axiom eqb_min_input_compl:∀h,x. + CF (λi.∑_{y ∈ [S i,S x[ }(h i y)) + (btotal (λi.eqb (min_input h i x) x)). *) +(*********************************** maximum **********************************) + +axiom CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s. + CF ha a → CF hb b → CF hp p → CF hf f → + O s (λx.ha x + hb x + ∑_{i ∈[a x ,b x[ }(hp 〈i,x〉 + hf 〈i,x〉)) → + CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). + +(******************************** minimization ********************************) + +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〉)). + +(****************************** constructibility ******************************) + +definition constructible ≝ λs. CF s s. + +lemma constr_comp : ∀s1,s2. constructible s1 → constructible s2 → + (∀x. x ≤ s2 x) → constructible (s2 ∘ s1). +#s1 #s2 #Hs1 #Hs2 #Hle @(CF_comp … Hs1 Hs2) @O_plus @le_to_O #x [@Hle | //] +qed. + +lemma ext_constr: ∀s1,s2. (∀x.s1 x = s2 x) → + constructible s1 → constructible s2. +#s1 #s2 #Hext #Hs1 @(ext_CF … Hext) @(monotonic_CF … Hs1) #x >Hext // +qed. + +(********************************* simulation *********************************) + +axiom sU : nat → nat. + +axiom monotonic_sU: ∀i1,i2,x1,x2,s1,s2. i1 ≤ i2 → x1 ≤ x2 → s1 ≤ s2 → + sU 〈i1,〈x1,s1〉〉 ≤ sU 〈i2,〈x2,s2〉〉. + +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) +#b1 * #c1 #eqy >eqy -eqy +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〉〉. + +definition pU_unary ≝ λp. pU (fst p) (fst (snd p)) (snd (snd p)). + +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) (termb_unary ∘ f). +#f @(CF_comp … CF_termb) *) + +(******************** complexity of g ********************) + +definition unary_g ≝ λh.λux. g h (fst ux) (snd ux). +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 (auxg h) → CF s (unary_g h). +#h #s #H1 @(CF_compS ? (auxg h) H1) +qed. + +definition aux1g ≝ + λh,ux. max_{i ∈[fst ux,snd ux[ | (λp. eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) 〈i,ux〉} + ((λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) 〈i,ux〉). + +lemma eq_aux : ∀h,x.aux1g h x = auxg h x. +#h #x @same_bigop + [#n #_ >fst_pair >snd_pair // |#n #_ #_ >fst_pair >snd_pair //] +qed. + +lemma compl_g2 : ∀h,s1,s2,s. + CF s1 + (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) → + CF s2 + (λ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 (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 (λ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 min_input_def whd in ⊢ (??%?); >minus_S_S @min_f_g #i #_ #_ +whd in ⊢ (??%%); >fst_pair >snd_pair // +qed. + +definition termb_aux ≝ λh. + termb_unary ∘ λp.〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉. + +(* +lemma termb_aux : ∀h,p. + (λx:ℕ.termb (fst x) (fst (snd x)) (snd (snd x))) + 〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉 = + termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)) . +#h #p normalize >fst_pair >snd_pair >fst_pair >snd_pair // +qed. *) + +lemma compl_g4 : ∀h,s1,s. + (CF s1 + (λ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 (λ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 + |@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //] +(* @(ext_CF (btotal (termb_aux h))) + [#n normalize >fst_pair >snd_pair >fst_pair >snd_pair // ] +@(CF_compb … CF_termb) *) +qed. + +(************************* a couple of technical lemmas ***********************) +lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0. +#a elim a // #n #Hind * + [#H @False_ind /2 by absurd/ | #b normalize #H @Hind @le_S_S_to_le /2/] +qed. + +lemma sigma_bound: ∀h,a,b. monotonic nat le h → + ∑_{i ∈ [a,S b[ }(h i) ≤ (S b-a)*h b. +#h #a #b #H cases (decidable_le a b) + [#leab cut (b = pred (S b - a + a)) + [Hb in match (h b); + generalize in match (S b -a); + #n elim n + [// + |#m #Hind >bigop_Strue [2://] @le_plus + [@H @le_n |@(transitive_le … Hind) @le_times [//] @H //] + ] + |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba + cut (S b -a = 0) [@minus_to_0 //] #Hcut >Hcut // + ] +qed. + +lemma sigma_bound_decr: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) → + ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a. +#h #a #b #H cases (decidable_le a b) + [#leab cut ((b -a) +a ≤ b) [/2 by le_minus_to_plus_r/] generalize in match (b -a); + #n elim n + [// + |#m #Hind >bigop_Strue [2://] #Hm + cut (m+a ≤ b) [@(transitive_le … Hm) //] #Hm1 + @le_plus [@H // |@(transitive_le … (Hind Hm1)) //] + ] + |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba + cut (b -a = 0) [@minus_to_0 @lt_to_le @ltba] #Hcut >Hcut // + ] +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〉)). +#s1 #Hs1 %{1} %{0} #n #_ >commutative_times minus_S_S //] +qed. + +lemma coroll2: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) → +O (λx.(snd x - fst x)*s1 〈fst x,x〉) (λx.∑_{i ∈[fst x,snd x[ }(s1 〈i,x〉)). +#s1 #Hs1 %{1} %{0} #n #_ >commutative_times fst_pair >snd_pair >fst_pair >snd_pair // ] +@(CF_comp … (λx.MSC 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 + [@O_plus + [@(O_trans … (λx.MSC (fst x) + MSC (max (fst (snd x)) (snd (snd x))))) + [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx + >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) + >distributive_times_plus @le_plus [//] + cases (surj_pair b) #c * #d #eqb >eqb + >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) + whd in ⊢ (??%); @le_plus + [@monotonic_MSC @(le_maxl … (le_n …)) + |>commutative_times fst_pair >snd_pair //] +@compl_g5 [2:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair +>fst_pair >snd_pair @monotonic_sU // @hmono @lexy +qed.*) + +definition big : nat →nat ≝ λx. + let m ≝ max (fst x) (snd x) in 〈m,m〉. + +lemma big_def : ∀a,b. big 〈a,b〉 = 〈max a b,max a b〉. +#a #b normalize >fst_pair >snd_pair // qed. + +lemma le_big : ∀x. x ≤ big x. +#x cases (surj_pair x) #a * #b #eqx >eqx @le_pair >fst_pair >snd_pair +[@(le_maxl … (le_n …)) | @(le_maxr … (le_n …))] +qed. + +definition faux2 ≝ λh. + (λx.MSC x + (snd (snd x)-fst x)* + (λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉). + +(* proviamo con x *) +lemma compl_g7: ∀h. + constructible (λx. h (fst x) (snd x)) → + (∀n. monotonic ? le (h n)) → + CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉) + (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #hcostr #hmono @(monotonic_CF … (faux2 h)) + [#n normalize >fst_pair >snd_pair //] +@compl_g5 [2:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair +>fst_pair >snd_pair @monotonic_sU // @hmono @lexy +qed. + +(* proviamo con x *) +lemma compl_g71: ∀h. + 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))〉〉) + (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #hcostr #hmono @(monotonic_CF … (compl_g7 h hcostr hmono)) #x +@le_plus [@monotonic_MSC //] +cases (decidable_le (fst x) (snd(snd x))) + [#Hle @le_times // @monotonic_sU + |#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt] + ] +qed. + +(* +axiom compl_g8: ∀h. + CF (λx. sU 〈fst 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)))). *) + +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 (MSC x)) >commutative_times @le_times + [// | @monotonic_MSC // ]] +@(O_trans … (coroll2 ??)) + [#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 [//|@(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 + [/2 by monotonic_le_minus_r/ + |@monotonic_sU // @hantimono [@le_S_S // |@ltb] + ] + |@monotonic_sU // @hantimono [@le_S_S // |@ltb] + ] + |@le_to_O #n >fst_pair >snd_pair + cut (max (fst n) n = n) [normalize >le_to_leb_true //] #Hmax >Hmax + >associative_plus >distributive_times_plus + @le_plus [@le_times [@le_S // |>big_def >Hmax //] |//] + ] +qed. + +definition sg ≝ λh,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)〉〉. + +lemma sg_def : ∀h,a,b. + sg h 〈a,b〉 = (S (b-a))*MSC 〈〈a,b〉,〈a,b〉〉 + + (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉. +#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair // +qed. + +lemma compl_g11 : ∀h. + 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) (unary_g h). +#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham) +qed. + +(**************************** closing the argument ****************************) + +let rec h_of_aux (r:nat →nat) (c,d,b:nat) on d : nat ≝ + match d with + [ O ⇒ c (* MSC 〈〈b,b〉,〈b,b〉〉 *) + | S d1 ⇒ (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + + d*(S d)*sU 〈〈b-d,b〉,〈b,r (h_of_aux r c d1 b)〉〉]. + +lemma h_of_aux_O: ∀r,c,b. + h_of_aux r c O b = c. +// qed. + +lemma h_of_aux_S : ∀r,c,d,b. + h_of_aux r c (S d) b = + (S (S d))*(MSC 〈〈b-(S d),b〉,〈b-(S d),b〉〉) + + (S d)*(S (S d))*sU 〈〈b-(S d),b〉,〈b,r(h_of_aux r c d b)〉〉. +// qed. + +definition h_of ≝ λr,p. + let m ≝ max (fst p) (snd p) in + h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (snd p - fst p) (snd p). + +lemma h_of_O: ∀r,a,b. b ≤ a → + h_of r 〈a,b〉 = let m ≝ max a b in MSC 〈〈m,m〉,〈m,m〉〉. +#r #a #b #Hle normalize >fst_pair >snd_pair >(minus_to_0 … Hle) // +qed. + +lemma h_of_def: ∀r,a,b.h_of r 〈a,b〉 = + let m ≝ max a b in + h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (b - a) b. +#r #a #b normalize >fst_pair >snd_pair // +qed. + +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. +#r #Hr #monor #d #d1 lapply d -d elim d1 + [#d #c #c1 #b #b1 #Hc #Hd @(le_n_O_elim ? Hd) #leb + >h_of_aux_O >h_of_aux_O // + |#m #Hind #d #c #c1 #b #b1 #lec #led #leb cases (le_to_or_lt_eq … led) + [#ltd @(transitive_le … (Hind … lec ? leb)) [@le_S_S_to_le @ltd] + >h_of_aux_S @(transitive_le ???? (le_plus_n …)) + >(times_n_1 (h_of_aux r c1 m b1)) in ⊢ (?%?); + >commutative_times @le_times [//|@(transitive_le … (Hr ?)) @sU_le] + |#Hd >Hd >h_of_aux_S >h_of_aux_S + cut (b-S m ≤ b1 - S m) [/2 by monotonic_le_minus_l/] #Hb1 + @le_plus [@le_times //] + [@monotonic_MSC @le_pair @le_pair // + |@le_times [//] @monotonic_sU + [@le_pair // |// |@monor @Hind //] + ] + ] + ] +qed. + +lemma mono_h_of2: ∀r.(∀x. x ≤ r x) → monotonic ? le r → + ∀i,b,b1. b ≤ b1 → h_of r 〈i,b〉 ≤ h_of r 〈i,b1〉. +#r #Hr #Hmono #i #a #b #leab >h_of_def >h_of_def +cut (max i a ≤ max i b) + [@to_max + [@(le_maxl … (le_n …))|@(transitive_le … leab) @(le_maxr … (le_n …))]] +#Hmax @(mono_h_of_aux r Hr Hmono) + [@monotonic_MSC @le_pair @le_pair @Hmax |/2 by monotonic_le_minus_l/ |@leab] +qed. + +axiom h_of_constr : ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + constructible (h_of r). + +lemma speed_compl: ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))). +#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …)) + [#x cases (surj_pair x) #a * #b #eqx >eqx + >sg_def cases (decidable_le b a) + [#leba >(minus_to_0 … leba) normalize in ⊢ (?%?); + h_of_def + cut (max a b = a) + [normalize cases (le_to_or_lt_eq … leba) + [#ltba >(lt_to_leb_false … ltba) % + |#eqba (le_to_leb_true … (le_n ?)) % ]] + #Hmax >Hmax normalize >(minus_to_0 … leba) normalize + @monotonic_MSC @le_pair @le_pair // + |#ltab >h_of_def >h_of_def + cut (max a b = b) + [normalize >(le_to_leb_true … ) [%] @lt_to_le @not_le_to_lt @ltab] + #Hmax >Hmax + 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) + [%{(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 + [@lt_to_le @not_le_to_lt // | //]] #eqd2 >eqd2 + normalize // + ] + |#n #a #b #leab #lebn >h_of_def >h_of_def + cut (max a n = n) + [normalize >le_to_leb_true [%|@(transitive_le … leab lebn)]] #Hmaxa + cut (max b n = n) + [normalize >(le_to_leb_true … lebn) %] #Hmaxb + >Hmaxa >Hmaxb @Hmono @(mono_h_of_aux r … Hr Hmono) // /2 by monotonic_le_minus_r/ + |#n #a #b #leab @Hmono @(mono_h_of2 … Hr Hmono … leab) + |@(constr_comp … Hconstr Hr) @(ext_constr (h_of r)) + [#x cases (surj_pair x) #a * #b #eqx >eqx >fst_pair >snd_pair //] + @(h_of_constr r Hr Hmono Hconstr) + ] +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. *) + +(* 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 → constructible r → + ∀i. CF (λx.h_of r 〈i,x〉) (λx.g (λi,x. r(h_of r 〈i,x〉)) i x). +#r #Hr #Hmono #Hconstr #i +@(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 Hconstr)) #n // +qed. + +theorem pseudo_speedup: + ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∃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 #Hconstr +(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * +#Hcodei #HCi +(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} +(* sg is (λx.h_of r 〈i,x〉) *) +%{(λx. h_of r 〈S i,x〉)} +lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg +%[%[@condition_1 |@Hg] + |cases Hg #H1 * #j * #Hcodej #HCj + lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) + cases HCi #m * #a #Ha %{m} %{(max (S i) a)} #n #ltin @lt_to_le @not_le_to_lt + @(not_to_not … Hcond2) -Hcond2 #Hlesf %{n} % + [@(transitive_le … ltin) @(le_maxl … (le_n …))] + cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] + #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) // + ] +qed. + +theorem pseudo_speedup': + ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∃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 #Hconstr +(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * +#Hcodei #HCi +(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} +(* sg is (λx.h_of r 〈i,x〉) *) +%{(λx. h_of r 〈S i,x〉)} +lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg +%[%[@condition_1 |@Hg] + |cases Hg #H1 * #j * #Hcodej #HCj + lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) + cases HCi #m * #a #Ha + %{m} %{(max (S i) a)} #n #ltin @not_le_to_lt @(not_to_not … Hcond2) -Hcond2 #Hlesf + %{n} % [@(transitive_le … ltin) @(le_maxl … (le_n …))] + cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] + #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) + @Hmono @(mono_h_of2 … Hr Hmono … ltin) + ] +qed. + diff --git a/matita/matita/lib/reverse_complexity/speed_def.ma b/matita/matita/lib/reverse_complexity/speed_def.ma new file mode 100644 index 000000000..9812cfb08 --- /dev/null +++ b/matita/matita/lib/reverse_complexity/speed_def.ma @@ -0,0 +1,922 @@ +include "basics/types.ma". +include "arithmetics/minimization.ma". +include "arithmetics/bigops.ma". +include "arithmetics/sigma_pi.ma". +include "arithmetics/bounded_quantifiers.ma". +include "reverse_complexity/big_O.ma". +include "basics/core_notation/napart_2.ma". + +(************************* notation for minimization *****************************) +notation "μ_{ ident i < n } p" + with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ≤ n } p" + with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b[ } p" + with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b] } p" + with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. + +(************************************ 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 le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true → + f a ≤ Max_{i ∈ [m,n[ | p i}(f i). +#f #p #n #m #a #lema #ltan #pa +>(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m)) + [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. + +(********************************** pairing ***********************************) +axiom pair: nat → nat → nat. +axiom fst : nat → nat. +axiom snd : nat → nat. + +interpretation "abstract pair" 'pair f g = (pair f g). + +axiom fst_pair: ∀a,b. fst 〈a,b〉 = a. +axiom snd_pair: ∀a,b. snd 〈a,b〉 = b. +axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉. + +axiom le_fst : ∀p. fst p ≤ p. +axiom le_snd : ∀p. snd p ≤ p. +axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉. + +(************************************* U **************************************) +axiom U: nat → nat →nat → option nat. + +axiom monotonic_U: ∀i,x,n,m,y.n ≤m → + U i x n = Some ? y → U i x m = Some ? y. + +lemma unique_U: ∀i,x,n,m,yn,ym. + U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. +#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. + +notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. + +lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n. +#i #x #n normalize cases (U i x n) + [%2 % * #y #H destruct|#y %1 %{y} //] +qed. + +lemma monotonic_terminate: ∀i,x,n,m. + n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m. +#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) // +qed. + +definition termb ≝ λi,x,t. + match U i x t with [None ⇒ false |Some y ⇒ true]. + +lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t. +#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] +qed. + +lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true. +#i #x #t * #y #H normalize >H // +qed. + +definition out ≝ λi,x,r. + match U i x r with [ None ⇒ 0 | Some z ⇒ z]. + +definition bool_to_nat: bool → nat ≝ + λb. match b with [true ⇒ 1 | false ⇒ 0]. + +coercion bool_to_nat. + +definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉. + +lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y. +#i #x #r #y % normalize + [cases (U i x r) normalize + [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H] + #H1 destruct + |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1] + #H1 // + ] + |#H >H //] +qed. + +lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?. +#i #x #r % normalize + [cases (U i x r) normalize // + #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1] + #H1 destruct + |#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)). + +lemma min_input_def : ∀h,i,x. + min_input h i x = min (x -i) (S i) (λy.termb i y (h (S i) y)). +// qed. + +lemma min_input_i: ∀h,i,x. x ≤ i → min_input h i x = S i. +#h #i #x #lexi >min_input_def +cut (x - i = 0) [@sym_eq /2 by eq_minus_O/] #Hcut // +qed. + +lemma min_input_to_terminate: ∀h,i,x. + min_input h i x = x → {i ⊙ x} ↓ (h (S i) x). +#h #i #x #Hminx +cases (decidable_le (S i) x) #Hix + [cases (true_or_false (termb i x (h (S i) x))) #Hcase + [@termb_true_to_term // + |min_input_def in Hminx; #Hminx >Hminx in ⊢ (%→?); + min_input_i in Hminx; + [#eqix >eqix in Hix; * /2/ | @le_S_S_to_le @not_le_to_lt //] + ] +qed. + +lemma min_input_to_lt: ∀h,i,x. + min_input h i x = x → i < x. +#h #i #x #Hminx cases (decidable_le (S i) x) // +#ltxi @False_ind >min_input_i in Hminx; + [#eqix >eqix in ltxi; * /2/ | @le_S_S_to_le @not_le_to_lt //] +qed. + +lemma le_to_min_input: ∀h,i,x,x1. x ≤ x1 → + min_input h i x = x → min_input h i x1 = x. +#h #i #x #x1 #lex #Hminx @(min_exists … (le_S_S … lex)) + [@(fmin_true … (sym_eq … Hminx)) // + |@(min_input_to_lt … Hminx) + |#j #H1 g_def cut (x-u = 0) [/2 by minus_le_minus_minus_comm/] +#eq0 >eq0 normalize // qed. + +lemma g_lt : ∀h,i,x. min_input h i x = x → + out i x (h (S i) x) < g h 0 x. +#h #i #x #H @le_S_S @(le_MaxI … i) /2 by min_input_to_lt/ +qed. + +lemma max_neq0 : ∀a,b. max a b ≠ 0 → a ≠ 0 ∨ b ≠ 0. +#a #b whd in match (max a b); cases (true_or_false (leb a b)) #Hcase >Hcase + [#H %2 @H | #H %1 @H] +qed. + +definition almost_equal ≝ λf,g:nat → nat. ¬ ∀nu.∃x. nu < x ∧ f x ≠ g x. +interpretation "almost equal" 'napart f g = (almost_equal f g). + +lemma eventually_cancelled: ∀h,u.¬∀nu.∃x. nu < x ∧ + max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) ≠ 0. +#h #u elim u + [normalize % #H cases (H u) #x * #_ * #H1 @H1 // + |#u0 @not_to_not #Hind #nu cases (Hind nu) #x * #ltx + cases (true_or_false (eqb (min_input h (u0+O) x) x)) #Hcase + [>bigop_Strue [2:@Hcase] #Hmax cases (max_neq0 … Hmax) -Hmax + [2: #H %{x} % // bigop_Sfalse + [#H %{x1} % [@transitive_lt //| (le_to_min_input … (eqb_true_to_eq … Hcase)) + [@lt_to_not_eq @ltx1 | @lt_to_le @ltx1] + ] + |>bigop_Sfalse [2:@Hcase] #H %{x} % // (bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA) + [>H // |@lt_to_le @(le_to_lt_to_lt …ltx) /2 by le_maxr/ |//] +qed. + +(******************************** Condition 2 *********************************) +definition total ≝ λf.λx:nat. Some nat (f x). + +lemma exists_to_exists_min: ∀h,i. (∃x. i < x ∧ {i ⊙ x} ↓ h (S i) x) → ∃y. min_input h i y = y. +#h #i * #x * #ltix #Hx %{(min_input h i x)} @min_spec_to_min @found // + [@(f_min_true (λy:ℕ.termb i y (h (S i) y))) %{x} % [% // | @term_to_termb_true //] + |#y #leiy #lty @(lt_min_to_false ????? lty) // + ] +qed. + +lemma condition_2: ∀h,i. code_for (total (g h 0)) i → ¬∃x. itermy >Hr +#H @(absurd ? H) @le_to_not_lt @le_n +qed. + + +(********************************* complexity *********************************) + +(* We assume operations have a minimal structural complexity MSC. +For instance, for time complexity, MSC is equal to the size of input. +For space complexity, MSC is typically 0, since we only measure the +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) *) + +definition C ≝ λs,i.∃c.∃a.∀x.a ≤ x → ∃y. + U i x (c*(s x)) = Some ? 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 ??); associative_times @le_times // @Ha1 @(transitive_le … lean) // + ] +qed. + +lemma timesc_CF: ∀s,f,c.CF (λx.c*s x) f → CF s f. +#s #f #c @O_to_CF @O_times_c +qed. + +(********************************* composition ********************************) +axiom CF_comp: ∀f,g,sf,sg,sh. CF sg g → CF sf f → + O sh (λx. sg x + sf (g x)) → CF sh (f ∘ g). + +lemma CF_comp_ext: ∀f,g,h,sh,sf,sg. CF sg g → CF sf f → + (∀x.f(g x) = h x) → O sh (λx. sg x + sf (g x)) → CF sh h. +#f #g #h #sh #sf #sg #Hg #Hf #Heq #H @(ext_CF (f ∘ g)) + [#n normalize @Heq | @(CF_comp … H) //] +qed. + + +(**************************** primitive operations*****************************) + +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_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 ***********************************) + +axiom CF_eqb: ∀h,f,g. + CF h f → CF h g → CF h (λx.eqb (f x) (g x)). + +(*********************************** maximum **********************************) + +axiom CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s. + CF ha a → CF hb b → CF hp p → CF hf f → + O s (λx.ha x + hb x + ∑_{i ∈[a x ,b x[ }(hp 〈i,x〉 + hf 〈i,x〉)) → + CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). + +(******************************** minimization ********************************) + +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〉)). + +(************************************* smn ************************************) +axiom smn: ∀f,s. CF s f → ∀x. CF (λy.s 〈x,y〉) (λy.f 〈x,y〉). + +(****************************** constructibility ******************************) + +definition constructible ≝ λs. CF s s. + +lemma constr_comp : ∀s1,s2. constructible s1 → constructible s2 → + (∀x. x ≤ s2 x) → constructible (s2 ∘ s1). +#s1 #s2 #Hs1 #Hs2 #Hle @(CF_comp … Hs1 Hs2) @O_plus @le_to_O #x [@Hle | //] +qed. + +lemma ext_constr: ∀s1,s2. (∀x.s1 x = s2 x) → + constructible s1 → constructible s2. +#s1 #s2 #Hext #Hs1 @(ext_CF … Hext) @(monotonic_CF … Hs1) #x >Hext // +qed. + +(********************************* simulation *********************************) + +axiom sU : nat → nat. + +axiom monotonic_sU: ∀i1,i2,x1,x2,s1,s2. i1 ≤ i2 → x1 ≤ x2 → s1 ≤ s2 → + sU 〈i1,〈x1,s1〉〉 ≤ sU 〈i2,〈x2,s2〉〉. + +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) +#b1 * #c1 #eqy >eqy -eqy +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〉〉. + +definition pU_unary ≝ λp. pU (fst p) (fst (snd p)) (snd (snd p)). + +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. + + +(******************** complexity of g ********************) + +definition unary_g ≝ λh.λux. g h (fst ux) (snd ux). +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 (auxg h) → CF s (unary_g h). +#h #s #H1 @(CF_compS ? (auxg h) H1) +qed. + +definition aux1g ≝ + λh,ux. max_{i ∈[fst ux,snd ux[ | (λp. eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) 〈i,ux〉} + ((λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) 〈i,ux〉). + +lemma eq_aux : ∀h,x.aux1g h x = auxg h x. +#h #x @same_bigop + [#n #_ >fst_pair >snd_pair // |#n #_ #_ >fst_pair >snd_pair //] +qed. + +lemma compl_g2 : ∀h,s1,s2,s. + CF s1 + (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) → + CF s2 + (λ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 (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 (λ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 min_input_def whd in ⊢ (??%?); >minus_S_S @min_f_g #i #_ #_ +whd in ⊢ (??%%); >fst_pair >snd_pair // +qed. + +definition termb_aux ≝ λh. + termb_unary ∘ λp.〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉. + +lemma compl_g4 : ∀h,s1,s. + (CF s1 + (λ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 (λ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 + |@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //] +qed. + +(************************* a couple of technical lemmas ***********************) +lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0. +#a elim a // #n #Hind * + [#H @False_ind /2 by absurd/ | #b normalize #H @Hind @le_S_S_to_le /2/] +qed. + +lemma sigma_bound: ∀h,a,b. monotonic nat le h → + ∑_{i ∈ [a,S b[ }(h i) ≤ (S b-a)*h b. +#h #a #b #H cases (decidable_le a b) + [#leab cut (b = pred (S b - a + a)) + [Hb in match (h b); + generalize in match (S b -a); + #n elim n + [// + |#m #Hind >bigop_Strue [2://] @le_plus + [@H @le_n |@(transitive_le … Hind) @le_times [//] @H //] + ] + |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba + cut (S b -a = 0) [@minus_to_0 //] #Hcut >Hcut // + ] +qed. + +lemma sigma_bound_decr: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) → + ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a. +#h #a #b #H cases (decidable_le a b) + [#leab cut ((b -a) +a ≤ b) [/2 by le_minus_to_plus_r/] generalize in match (b -a); + #n elim n + [// + |#m #Hind >bigop_Strue [2://] #Hm + cut (m+a ≤ b) [@(transitive_le … Hm) //] #Hm1 + @le_plus [@H // |@(transitive_le … (Hind Hm1)) //] + ] + |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba + cut (b -a = 0) [@minus_to_0 @lt_to_le @ltba] #Hcut >Hcut // + ] +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〉)). +#s1 #Hs1 %{1} %{0} #n #_ >commutative_times minus_S_S //] +qed. + +lemma coroll2: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) → +O (λx.(snd x - fst x)*s1 〈fst x,x〉) (λx.∑_{i ∈[fst x,snd x[ }(s1 〈i,x〉)). +#s1 #Hs1 %{1} %{0} #n #_ >commutative_times fst_pair >snd_pair >fst_pair >snd_pair // ] +@(CF_comp … (λx.MSC 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 + [@O_plus + [@(O_trans … (λx.MSC (fst x) + MSC (max (fst (snd x)) (snd (snd x))))) + [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx + >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) + >distributive_times_plus @le_plus [//] + cases (surj_pair b) #c * #d #eqb >eqb + >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) + whd in ⊢ (??%); @le_plus + [@monotonic_MSC @(le_maxl … (le_n …)) + |>commutative_times fst_pair >snd_pair // qed. + +lemma le_big : ∀x. x ≤ big x. +#x cases (surj_pair x) #a * #b #eqx >eqx @le_pair >fst_pair >snd_pair +[@(le_maxl … (le_n …)) | @(le_maxr … (le_n …))] +qed. + +definition faux2 ≝ λh. + (λx.MSC x + (snd (snd x)-fst x)* + (λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉). + +lemma compl_g7: ∀h. + constructible (λx. h (fst x) (snd x)) → + (∀n. monotonic ? le (h n)) → + CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉) + (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #hcostr #hmono @(monotonic_CF … (faux2 h)) + [#n normalize >fst_pair >snd_pair //] +@compl_g5 [2:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair +>fst_pair >snd_pair @monotonic_sU // @hmono @lexy +qed. + +lemma compl_g71: ∀h. + 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))〉〉) + (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #hcostr #hmono @(monotonic_CF … (compl_g7 h hcostr hmono)) #x +@le_plus [@monotonic_MSC //] +cases (decidable_le (fst x) (snd(snd x))) + [#Hle @le_times // @monotonic_sU + |#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt] + ] +qed. + +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 (MSC x)) >commutative_times @le_times + [// | @monotonic_MSC // ]] +@(O_trans … (coroll2 ??)) + [#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 [//|@(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 + [/2 by monotonic_le_minus_r/ + |@monotonic_sU // @hantimono [@le_S_S // |@ltb] + ] + |@monotonic_sU // @hantimono [@le_S_S // |@ltb] + ] + |@le_to_O #n >fst_pair >snd_pair + cut (max (fst n) n = n) [normalize >le_to_leb_true //] #Hmax >Hmax + >associative_plus >distributive_times_plus + @le_plus [@le_times [@le_S // |>big_def >Hmax //] |//] + ] +qed. + +definition sg ≝ λh,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)〉〉. + +lemma sg_def : ∀h,a,b. + sg h 〈a,b〉 = (S (b-a))*MSC 〈〈a,b〉,〈a,b〉〉 + + (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉. +#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair // +qed. + +lemma compl_g11 : ∀h. + 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) (unary_g h). +#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham) +qed. + +(**************************** closing the argument ****************************) + +let rec h_of_aux (r:nat →nat) (c,d,b:nat) on d : nat ≝ + match d with + [ O ⇒ c + | S d1 ⇒ (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + + d*(S d)*sU 〈〈b-d,b〉,〈b,r (h_of_aux r c d1 b)〉〉]. + +lemma h_of_aux_O: ∀r,c,b. + h_of_aux r c O b = c. +// qed. + +lemma h_of_aux_S : ∀r,c,d,b. + h_of_aux r c (S d) b = + (S (S d))*(MSC 〈〈b-(S d),b〉,〈b-(S d),b〉〉) + + (S d)*(S (S d))*sU 〈〈b-(S d),b〉,〈b,r(h_of_aux r c d b)〉〉. +// qed. + +definition h_of ≝ λr,p. + let m ≝ max (fst p) (snd p) in + h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (snd p - fst p) (snd p). + +lemma h_of_O: ∀r,a,b. b ≤ a → + h_of r 〈a,b〉 = let m ≝ max a b in MSC 〈〈m,m〉,〈m,m〉〉. +#r #a #b #Hle normalize >fst_pair >snd_pair >(minus_to_0 … Hle) // +qed. + +lemma h_of_def: ∀r,a,b.h_of r 〈a,b〉 = + let m ≝ max a b in + h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (b - a) b. +#r #a #b normalize >fst_pair >snd_pair // +qed. + +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. +#r #Hr #monor #d #d1 lapply d -d elim d1 + [#d #c #c1 #b #b1 #Hc #Hd @(le_n_O_elim ? Hd) #leb + >h_of_aux_O >h_of_aux_O // + |#m #Hind #d #c #c1 #b #b1 #lec #led #leb cases (le_to_or_lt_eq … led) + [#ltd @(transitive_le … (Hind … lec ? leb)) [@le_S_S_to_le @ltd] + >h_of_aux_S @(transitive_le ???? (le_plus_n …)) + >(times_n_1 (h_of_aux r c1 m b1)) in ⊢ (?%?); + >commutative_times @le_times [//|@(transitive_le … (Hr ?)) @sU_le] + |#Hd >Hd >h_of_aux_S >h_of_aux_S + cut (b-S m ≤ b1 - S m) [/2 by monotonic_le_minus_l/] #Hb1 + @le_plus [@le_times //] + [@monotonic_MSC @le_pair @le_pair // + |@le_times [//] @monotonic_sU + [@le_pair // |// |@monor @Hind //] + ] + ] + ] +qed. + +lemma mono_h_of2: ∀r.(∀x. x ≤ r x) → monotonic ? le r → + ∀i,b,b1. b ≤ b1 → h_of r 〈i,b〉 ≤ h_of r 〈i,b1〉. +#r #Hr #Hmono #i #a #b #leab >h_of_def >h_of_def +cut (max i a ≤ max i b) + [@to_max + [@(le_maxl … (le_n …))|@(transitive_le … leab) @(le_maxr … (le_n …))]] +#Hmax @(mono_h_of_aux r Hr Hmono) + [@monotonic_MSC @le_pair @le_pair @Hmax |/2 by monotonic_le_minus_l/ |@leab] +qed. + +axiom h_of_constr : ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + constructible (h_of r). + +lemma speed_compl: ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))). +#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …)) + [#x cases (surj_pair x) #a * #b #eqx >eqx + >sg_def cases (decidable_le b a) + [#leba >(minus_to_0 … leba) normalize in ⊢ (?%?); + h_of_def + cut (max a b = a) + [normalize cases (le_to_or_lt_eq … leba) + [#ltba >(lt_to_leb_false … ltba) % + |#eqba (le_to_leb_true … (le_n ?)) % ]] + #Hmax >Hmax normalize >(minus_to_0 … leba) normalize + @monotonic_MSC @le_pair @le_pair // + |#ltab >h_of_def >h_of_def + cut (max a b = b) + [normalize >(le_to_leb_true … ) [%] @lt_to_le @not_le_to_lt @ltab] + #Hmax >Hmax + 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) + [%{(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 + [@lt_to_le @not_le_to_lt // | //]] #eqd2 >eqd2 + normalize // + ] + |#n #a #b #leab #lebn >h_of_def >h_of_def + cut (max a n = n) + [normalize >le_to_leb_true [%|@(transitive_le … leab lebn)]] #Hmaxa + cut (max b n = n) + [normalize >(le_to_leb_true … lebn) %] #Hmaxb + >Hmaxa >Hmaxb @Hmono @(mono_h_of_aux r … Hr Hmono) // /2 by monotonic_le_minus_r/ + |#n #a #b #leab @Hmono @(mono_h_of2 … Hr Hmono … leab) + |@(constr_comp … Hconstr Hr) @(ext_constr (h_of r)) + [#x cases (surj_pair x) #a * #b #eqx >eqx >fst_pair >snd_pair //] + @(h_of_constr r Hr Hmono Hconstr) + ] +qed. + +lemma speed_compl_i: ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∀i. CF (λx.h_of r 〈i,x〉) (λx.g (λi,x. r(h_of r 〈i,x〉)) i x). +#r #Hr #Hmono #Hconstr #i +@(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 Hconstr)) #n // +qed. + +(**************************** the speedup theorem *****************************) +theorem pseudo_speedup: + ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∃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 #Hconstr +(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * +#Hcodei #HCi +(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} +(* sg is (λx.h_of r 〈i,x〉) *) +%{(λx. h_of r 〈S i,x〉)} +lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg +%[%[@condition_1 |@Hg] + |cases Hg #H1 * #j * #Hcodej #HCj + lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) + cases HCi #m * #a #Ha %{m} %{(max (S i) a)} #n #ltin @lt_to_le @not_le_to_lt + @(not_to_not … Hcond2) -Hcond2 #Hlesf %{n} % + [@(transitive_le … ltin) @(le_maxl … (le_n …))] + cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] + #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) // + ] +qed. + +theorem pseudo_speedup': + ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∃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 #Hconstr +(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * +#Hcodei #HCi +(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} +(* sg is (λx.h_of r 〈i,x〉) *) +%{(λx. h_of r 〈S i,x〉)} +lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg +%[%[@condition_1 |@Hg] + |cases Hg #H1 * #j * #Hcodej #HCj + lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) + cases HCi #m * #a #Ha + %{m} %{(max (S i) a)} #n #ltin @not_le_to_lt @(not_to_not … Hcond2) -Hcond2 #Hlesf + %{n} % [@(transitive_le … ltin) @(le_maxl … (le_n …))] + cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] + #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) + @Hmono @(mono_h_of2 … Hr Hmono … ltin) + ] +qed. + diff --git a/matita/matita/lib/reverse_complexity/speed_new.ma b/matita/matita/lib/reverse_complexity/speed_new.ma new file mode 100644 index 000000000..9812cfb08 --- /dev/null +++ b/matita/matita/lib/reverse_complexity/speed_new.ma @@ -0,0 +1,922 @@ +include "basics/types.ma". +include "arithmetics/minimization.ma". +include "arithmetics/bigops.ma". +include "arithmetics/sigma_pi.ma". +include "arithmetics/bounded_quantifiers.ma". +include "reverse_complexity/big_O.ma". +include "basics/core_notation/napart_2.ma". + +(************************* notation for minimization *****************************) +notation "μ_{ ident i < n } p" + with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ≤ n } p" + with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b[ } p" + with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b] } p" + with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. + +(************************************ 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 le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true → + f a ≤ Max_{i ∈ [m,n[ | p i}(f i). +#f #p #n #m #a #lema #ltan #pa +>(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m)) + [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. + +(********************************** pairing ***********************************) +axiom pair: nat → nat → nat. +axiom fst : nat → nat. +axiom snd : nat → nat. + +interpretation "abstract pair" 'pair f g = (pair f g). + +axiom fst_pair: ∀a,b. fst 〈a,b〉 = a. +axiom snd_pair: ∀a,b. snd 〈a,b〉 = b. +axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉. + +axiom le_fst : ∀p. fst p ≤ p. +axiom le_snd : ∀p. snd p ≤ p. +axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉. + +(************************************* U **************************************) +axiom U: nat → nat →nat → option nat. + +axiom monotonic_U: ∀i,x,n,m,y.n ≤m → + U i x n = Some ? y → U i x m = Some ? y. + +lemma unique_U: ∀i,x,n,m,yn,ym. + U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. +#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. + +notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. + +lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n. +#i #x #n normalize cases (U i x n) + [%2 % * #y #H destruct|#y %1 %{y} //] +qed. + +lemma monotonic_terminate: ∀i,x,n,m. + n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m. +#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) // +qed. + +definition termb ≝ λi,x,t. + match U i x t with [None ⇒ false |Some y ⇒ true]. + +lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t. +#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] +qed. + +lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true. +#i #x #t * #y #H normalize >H // +qed. + +definition out ≝ λi,x,r. + match U i x r with [ None ⇒ 0 | Some z ⇒ z]. + +definition bool_to_nat: bool → nat ≝ + λb. match b with [true ⇒ 1 | false ⇒ 0]. + +coercion bool_to_nat. + +definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉. + +lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y. +#i #x #r #y % normalize + [cases (U i x r) normalize + [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H] + #H1 destruct + |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1] + #H1 // + ] + |#H >H //] +qed. + +lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?. +#i #x #r % normalize + [cases (U i x r) normalize // + #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1] + #H1 destruct + |#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)). + +lemma min_input_def : ∀h,i,x. + min_input h i x = min (x -i) (S i) (λy.termb i y (h (S i) y)). +// qed. + +lemma min_input_i: ∀h,i,x. x ≤ i → min_input h i x = S i. +#h #i #x #lexi >min_input_def +cut (x - i = 0) [@sym_eq /2 by eq_minus_O/] #Hcut // +qed. + +lemma min_input_to_terminate: ∀h,i,x. + min_input h i x = x → {i ⊙ x} ↓ (h (S i) x). +#h #i #x #Hminx +cases (decidable_le (S i) x) #Hix + [cases (true_or_false (termb i x (h (S i) x))) #Hcase + [@termb_true_to_term // + |min_input_def in Hminx; #Hminx >Hminx in ⊢ (%→?); + min_input_i in Hminx; + [#eqix >eqix in Hix; * /2/ | @le_S_S_to_le @not_le_to_lt //] + ] +qed. + +lemma min_input_to_lt: ∀h,i,x. + min_input h i x = x → i < x. +#h #i #x #Hminx cases (decidable_le (S i) x) // +#ltxi @False_ind >min_input_i in Hminx; + [#eqix >eqix in ltxi; * /2/ | @le_S_S_to_le @not_le_to_lt //] +qed. + +lemma le_to_min_input: ∀h,i,x,x1. x ≤ x1 → + min_input h i x = x → min_input h i x1 = x. +#h #i #x #x1 #lex #Hminx @(min_exists … (le_S_S … lex)) + [@(fmin_true … (sym_eq … Hminx)) // + |@(min_input_to_lt … Hminx) + |#j #H1 g_def cut (x-u = 0) [/2 by minus_le_minus_minus_comm/] +#eq0 >eq0 normalize // qed. + +lemma g_lt : ∀h,i,x. min_input h i x = x → + out i x (h (S i) x) < g h 0 x. +#h #i #x #H @le_S_S @(le_MaxI … i) /2 by min_input_to_lt/ +qed. + +lemma max_neq0 : ∀a,b. max a b ≠ 0 → a ≠ 0 ∨ b ≠ 0. +#a #b whd in match (max a b); cases (true_or_false (leb a b)) #Hcase >Hcase + [#H %2 @H | #H %1 @H] +qed. + +definition almost_equal ≝ λf,g:nat → nat. ¬ ∀nu.∃x. nu < x ∧ f x ≠ g x. +interpretation "almost equal" 'napart f g = (almost_equal f g). + +lemma eventually_cancelled: ∀h,u.¬∀nu.∃x. nu < x ∧ + max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) ≠ 0. +#h #u elim u + [normalize % #H cases (H u) #x * #_ * #H1 @H1 // + |#u0 @not_to_not #Hind #nu cases (Hind nu) #x * #ltx + cases (true_or_false (eqb (min_input h (u0+O) x) x)) #Hcase + [>bigop_Strue [2:@Hcase] #Hmax cases (max_neq0 … Hmax) -Hmax + [2: #H %{x} % // bigop_Sfalse + [#H %{x1} % [@transitive_lt //| (le_to_min_input … (eqb_true_to_eq … Hcase)) + [@lt_to_not_eq @ltx1 | @lt_to_le @ltx1] + ] + |>bigop_Sfalse [2:@Hcase] #H %{x} % // (bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA) + [>H // |@lt_to_le @(le_to_lt_to_lt …ltx) /2 by le_maxr/ |//] +qed. + +(******************************** Condition 2 *********************************) +definition total ≝ λf.λx:nat. Some nat (f x). + +lemma exists_to_exists_min: ∀h,i. (∃x. i < x ∧ {i ⊙ x} ↓ h (S i) x) → ∃y. min_input h i y = y. +#h #i * #x * #ltix #Hx %{(min_input h i x)} @min_spec_to_min @found // + [@(f_min_true (λy:ℕ.termb i y (h (S i) y))) %{x} % [% // | @term_to_termb_true //] + |#y #leiy #lty @(lt_min_to_false ????? lty) // + ] +qed. + +lemma condition_2: ∀h,i. code_for (total (g h 0)) i → ¬∃x. itermy >Hr +#H @(absurd ? H) @le_to_not_lt @le_n +qed. + + +(********************************* complexity *********************************) + +(* We assume operations have a minimal structural complexity MSC. +For instance, for time complexity, MSC is equal to the size of input. +For space complexity, MSC is typically 0, since we only measure the +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) *) + +definition C ≝ λs,i.∃c.∃a.∀x.a ≤ x → ∃y. + U i x (c*(s x)) = Some ? 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 ??); associative_times @le_times // @Ha1 @(transitive_le … lean) // + ] +qed. + +lemma timesc_CF: ∀s,f,c.CF (λx.c*s x) f → CF s f. +#s #f #c @O_to_CF @O_times_c +qed. + +(********************************* composition ********************************) +axiom CF_comp: ∀f,g,sf,sg,sh. CF sg g → CF sf f → + O sh (λx. sg x + sf (g x)) → CF sh (f ∘ g). + +lemma CF_comp_ext: ∀f,g,h,sh,sf,sg. CF sg g → CF sf f → + (∀x.f(g x) = h x) → O sh (λx. sg x + sf (g x)) → CF sh h. +#f #g #h #sh #sf #sg #Hg #Hf #Heq #H @(ext_CF (f ∘ g)) + [#n normalize @Heq | @(CF_comp … H) //] +qed. + + +(**************************** primitive operations*****************************) + +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_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 ***********************************) + +axiom CF_eqb: ∀h,f,g. + CF h f → CF h g → CF h (λx.eqb (f x) (g x)). + +(*********************************** maximum **********************************) + +axiom CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s. + CF ha a → CF hb b → CF hp p → CF hf f → + O s (λx.ha x + hb x + ∑_{i ∈[a x ,b x[ }(hp 〈i,x〉 + hf 〈i,x〉)) → + CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). + +(******************************** minimization ********************************) + +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〉)). + +(************************************* smn ************************************) +axiom smn: ∀f,s. CF s f → ∀x. CF (λy.s 〈x,y〉) (λy.f 〈x,y〉). + +(****************************** constructibility ******************************) + +definition constructible ≝ λs. CF s s. + +lemma constr_comp : ∀s1,s2. constructible s1 → constructible s2 → + (∀x. x ≤ s2 x) → constructible (s2 ∘ s1). +#s1 #s2 #Hs1 #Hs2 #Hle @(CF_comp … Hs1 Hs2) @O_plus @le_to_O #x [@Hle | //] +qed. + +lemma ext_constr: ∀s1,s2. (∀x.s1 x = s2 x) → + constructible s1 → constructible s2. +#s1 #s2 #Hext #Hs1 @(ext_CF … Hext) @(monotonic_CF … Hs1) #x >Hext // +qed. + +(********************************* simulation *********************************) + +axiom sU : nat → nat. + +axiom monotonic_sU: ∀i1,i2,x1,x2,s1,s2. i1 ≤ i2 → x1 ≤ x2 → s1 ≤ s2 → + sU 〈i1,〈x1,s1〉〉 ≤ sU 〈i2,〈x2,s2〉〉. + +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) +#b1 * #c1 #eqy >eqy -eqy +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〉〉. + +definition pU_unary ≝ λp. pU (fst p) (fst (snd p)) (snd (snd p)). + +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. + + +(******************** complexity of g ********************) + +definition unary_g ≝ λh.λux. g h (fst ux) (snd ux). +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 (auxg h) → CF s (unary_g h). +#h #s #H1 @(CF_compS ? (auxg h) H1) +qed. + +definition aux1g ≝ + λh,ux. max_{i ∈[fst ux,snd ux[ | (λp. eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) 〈i,ux〉} + ((λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) 〈i,ux〉). + +lemma eq_aux : ∀h,x.aux1g h x = auxg h x. +#h #x @same_bigop + [#n #_ >fst_pair >snd_pair // |#n #_ #_ >fst_pair >snd_pair //] +qed. + +lemma compl_g2 : ∀h,s1,s2,s. + CF s1 + (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) → + CF s2 + (λ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 (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 (λ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 min_input_def whd in ⊢ (??%?); >minus_S_S @min_f_g #i #_ #_ +whd in ⊢ (??%%); >fst_pair >snd_pair // +qed. + +definition termb_aux ≝ λh. + termb_unary ∘ λp.〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉. + +lemma compl_g4 : ∀h,s1,s. + (CF s1 + (λ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 (λ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 + |@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //] +qed. + +(************************* a couple of technical lemmas ***********************) +lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0. +#a elim a // #n #Hind * + [#H @False_ind /2 by absurd/ | #b normalize #H @Hind @le_S_S_to_le /2/] +qed. + +lemma sigma_bound: ∀h,a,b. monotonic nat le h → + ∑_{i ∈ [a,S b[ }(h i) ≤ (S b-a)*h b. +#h #a #b #H cases (decidable_le a b) + [#leab cut (b = pred (S b - a + a)) + [Hb in match (h b); + generalize in match (S b -a); + #n elim n + [// + |#m #Hind >bigop_Strue [2://] @le_plus + [@H @le_n |@(transitive_le … Hind) @le_times [//] @H //] + ] + |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba + cut (S b -a = 0) [@minus_to_0 //] #Hcut >Hcut // + ] +qed. + +lemma sigma_bound_decr: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) → + ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a. +#h #a #b #H cases (decidable_le a b) + [#leab cut ((b -a) +a ≤ b) [/2 by le_minus_to_plus_r/] generalize in match (b -a); + #n elim n + [// + |#m #Hind >bigop_Strue [2://] #Hm + cut (m+a ≤ b) [@(transitive_le … Hm) //] #Hm1 + @le_plus [@H // |@(transitive_le … (Hind Hm1)) //] + ] + |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba + cut (b -a = 0) [@minus_to_0 @lt_to_le @ltba] #Hcut >Hcut // + ] +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〉)). +#s1 #Hs1 %{1} %{0} #n #_ >commutative_times minus_S_S //] +qed. + +lemma coroll2: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) → +O (λx.(snd x - fst x)*s1 〈fst x,x〉) (λx.∑_{i ∈[fst x,snd x[ }(s1 〈i,x〉)). +#s1 #Hs1 %{1} %{0} #n #_ >commutative_times fst_pair >snd_pair >fst_pair >snd_pair // ] +@(CF_comp … (λx.MSC 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 + [@O_plus + [@(O_trans … (λx.MSC (fst x) + MSC (max (fst (snd x)) (snd (snd x))))) + [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx + >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) + >distributive_times_plus @le_plus [//] + cases (surj_pair b) #c * #d #eqb >eqb + >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) + whd in ⊢ (??%); @le_plus + [@monotonic_MSC @(le_maxl … (le_n …)) + |>commutative_times fst_pair >snd_pair // qed. + +lemma le_big : ∀x. x ≤ big x. +#x cases (surj_pair x) #a * #b #eqx >eqx @le_pair >fst_pair >snd_pair +[@(le_maxl … (le_n …)) | @(le_maxr … (le_n …))] +qed. + +definition faux2 ≝ λh. + (λx.MSC x + (snd (snd x)-fst x)* + (λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉). + +lemma compl_g7: ∀h. + constructible (λx. h (fst x) (snd x)) → + (∀n. monotonic ? le (h n)) → + CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉) + (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #hcostr #hmono @(monotonic_CF … (faux2 h)) + [#n normalize >fst_pair >snd_pair //] +@compl_g5 [2:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair +>fst_pair >snd_pair @monotonic_sU // @hmono @lexy +qed. + +lemma compl_g71: ∀h. + 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))〉〉) + (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #hcostr #hmono @(monotonic_CF … (compl_g7 h hcostr hmono)) #x +@le_plus [@monotonic_MSC //] +cases (decidable_le (fst x) (snd(snd x))) + [#Hle @le_times // @monotonic_sU + |#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt] + ] +qed. + +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 (MSC x)) >commutative_times @le_times + [// | @monotonic_MSC // ]] +@(O_trans … (coroll2 ??)) + [#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 [//|@(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 + [/2 by monotonic_le_minus_r/ + |@monotonic_sU // @hantimono [@le_S_S // |@ltb] + ] + |@monotonic_sU // @hantimono [@le_S_S // |@ltb] + ] + |@le_to_O #n >fst_pair >snd_pair + cut (max (fst n) n = n) [normalize >le_to_leb_true //] #Hmax >Hmax + >associative_plus >distributive_times_plus + @le_plus [@le_times [@le_S // |>big_def >Hmax //] |//] + ] +qed. + +definition sg ≝ λh,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)〉〉. + +lemma sg_def : ∀h,a,b. + sg h 〈a,b〉 = (S (b-a))*MSC 〈〈a,b〉,〈a,b〉〉 + + (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉. +#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair // +qed. + +lemma compl_g11 : ∀h. + 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) (unary_g h). +#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham) +qed. + +(**************************** closing the argument ****************************) + +let rec h_of_aux (r:nat →nat) (c,d,b:nat) on d : nat ≝ + match d with + [ O ⇒ c + | S d1 ⇒ (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + + d*(S d)*sU 〈〈b-d,b〉,〈b,r (h_of_aux r c d1 b)〉〉]. + +lemma h_of_aux_O: ∀r,c,b. + h_of_aux r c O b = c. +// qed. + +lemma h_of_aux_S : ∀r,c,d,b. + h_of_aux r c (S d) b = + (S (S d))*(MSC 〈〈b-(S d),b〉,〈b-(S d),b〉〉) + + (S d)*(S (S d))*sU 〈〈b-(S d),b〉,〈b,r(h_of_aux r c d b)〉〉. +// qed. + +definition h_of ≝ λr,p. + let m ≝ max (fst p) (snd p) in + h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (snd p - fst p) (snd p). + +lemma h_of_O: ∀r,a,b. b ≤ a → + h_of r 〈a,b〉 = let m ≝ max a b in MSC 〈〈m,m〉,〈m,m〉〉. +#r #a #b #Hle normalize >fst_pair >snd_pair >(minus_to_0 … Hle) // +qed. + +lemma h_of_def: ∀r,a,b.h_of r 〈a,b〉 = + let m ≝ max a b in + h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (b - a) b. +#r #a #b normalize >fst_pair >snd_pair // +qed. + +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. +#r #Hr #monor #d #d1 lapply d -d elim d1 + [#d #c #c1 #b #b1 #Hc #Hd @(le_n_O_elim ? Hd) #leb + >h_of_aux_O >h_of_aux_O // + |#m #Hind #d #c #c1 #b #b1 #lec #led #leb cases (le_to_or_lt_eq … led) + [#ltd @(transitive_le … (Hind … lec ? leb)) [@le_S_S_to_le @ltd] + >h_of_aux_S @(transitive_le ???? (le_plus_n …)) + >(times_n_1 (h_of_aux r c1 m b1)) in ⊢ (?%?); + >commutative_times @le_times [//|@(transitive_le … (Hr ?)) @sU_le] + |#Hd >Hd >h_of_aux_S >h_of_aux_S + cut (b-S m ≤ b1 - S m) [/2 by monotonic_le_minus_l/] #Hb1 + @le_plus [@le_times //] + [@monotonic_MSC @le_pair @le_pair // + |@le_times [//] @monotonic_sU + [@le_pair // |// |@monor @Hind //] + ] + ] + ] +qed. + +lemma mono_h_of2: ∀r.(∀x. x ≤ r x) → monotonic ? le r → + ∀i,b,b1. b ≤ b1 → h_of r 〈i,b〉 ≤ h_of r 〈i,b1〉. +#r #Hr #Hmono #i #a #b #leab >h_of_def >h_of_def +cut (max i a ≤ max i b) + [@to_max + [@(le_maxl … (le_n …))|@(transitive_le … leab) @(le_maxr … (le_n …))]] +#Hmax @(mono_h_of_aux r Hr Hmono) + [@monotonic_MSC @le_pair @le_pair @Hmax |/2 by monotonic_le_minus_l/ |@leab] +qed. + +axiom h_of_constr : ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + constructible (h_of r). + +lemma speed_compl: ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))). +#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …)) + [#x cases (surj_pair x) #a * #b #eqx >eqx + >sg_def cases (decidable_le b a) + [#leba >(minus_to_0 … leba) normalize in ⊢ (?%?); + h_of_def + cut (max a b = a) + [normalize cases (le_to_or_lt_eq … leba) + [#ltba >(lt_to_leb_false … ltba) % + |#eqba (le_to_leb_true … (le_n ?)) % ]] + #Hmax >Hmax normalize >(minus_to_0 … leba) normalize + @monotonic_MSC @le_pair @le_pair // + |#ltab >h_of_def >h_of_def + cut (max a b = b) + [normalize >(le_to_leb_true … ) [%] @lt_to_le @not_le_to_lt @ltab] + #Hmax >Hmax + 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) + [%{(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 + [@lt_to_le @not_le_to_lt // | //]] #eqd2 >eqd2 + normalize // + ] + |#n #a #b #leab #lebn >h_of_def >h_of_def + cut (max a n = n) + [normalize >le_to_leb_true [%|@(transitive_le … leab lebn)]] #Hmaxa + cut (max b n = n) + [normalize >(le_to_leb_true … lebn) %] #Hmaxb + >Hmaxa >Hmaxb @Hmono @(mono_h_of_aux r … Hr Hmono) // /2 by monotonic_le_minus_r/ + |#n #a #b #leab @Hmono @(mono_h_of2 … Hr Hmono … leab) + |@(constr_comp … Hconstr Hr) @(ext_constr (h_of r)) + [#x cases (surj_pair x) #a * #b #eqx >eqx >fst_pair >snd_pair //] + @(h_of_constr r Hr Hmono Hconstr) + ] +qed. + +lemma speed_compl_i: ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∀i. CF (λx.h_of r 〈i,x〉) (λx.g (λi,x. r(h_of r 〈i,x〉)) i x). +#r #Hr #Hmono #Hconstr #i +@(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 Hconstr)) #n // +qed. + +(**************************** the speedup theorem *****************************) +theorem pseudo_speedup: + ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∃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 #Hconstr +(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * +#Hcodei #HCi +(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} +(* sg is (λx.h_of r 〈i,x〉) *) +%{(λx. h_of r 〈S i,x〉)} +lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg +%[%[@condition_1 |@Hg] + |cases Hg #H1 * #j * #Hcodej #HCj + lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) + cases HCi #m * #a #Ha %{m} %{(max (S i) a)} #n #ltin @lt_to_le @not_le_to_lt + @(not_to_not … Hcond2) -Hcond2 #Hlesf %{n} % + [@(transitive_le … ltin) @(le_maxl … (le_n …))] + cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] + #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) // + ] +qed. + +theorem pseudo_speedup': + ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∃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 #Hconstr +(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * +#Hcodei #HCi +(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} +(* sg is (λx.h_of r 〈i,x〉) *) +%{(λx. h_of r 〈S i,x〉)} +lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg +%[%[@condition_1 |@Hg] + |cases Hg #H1 * #j * #Hcodej #HCj + lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) + cases HCi #m * #a #Ha + %{m} %{(max (S i) a)} #n #ltin @not_le_to_lt @(not_to_not … Hcond2) -Hcond2 #Hlesf + %{n} % [@(transitive_le … ltin) @(le_maxl … (le_n …))] + cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] + #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) + @Hmono @(mono_h_of2 … Hr Hmono … ltin) + ] +qed. + diff --git a/matita/matita/lib/reverse_complexity/speedup.ma b/matita/matita/lib/reverse_complexity/speedup.ma new file mode 100644 index 000000000..12605c9e4 --- /dev/null +++ b/matita/matita/lib/reverse_complexity/speedup.ma @@ -0,0 +1,760 @@ +(* +<<<<<<< HEAD:matita/matita/broken_lib/reverse_complexity/toolkit.ma +include "basics/types.ma". +include "arithmetics/minimization.ma". +include "arithmetics/bigops.ma". +include "arithmetics/sigma_pi.ma". +include "arithmetics/bounded_quantifiers.ma". +include "reverse_complexity/big_O.ma". +include "basics/core_notation/napart_2.ma". + +(************************* notation for minimization *****************************) +notation "μ_{ ident i < n } p" + with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ≤ n } p" + with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b[ } p" + with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b] } p" + with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. + +(************************************ 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 le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true → + f a ≤ Max_{i ∈ [m,n[ | p i}(f i). +#f #p #n #m #a #lema #ltan #pa +>(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m)) + [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. + +(********************************** pairing ***********************************) +axiom pair: nat → nat → nat. +axiom fst : nat → nat. +axiom snd : nat → nat. + +interpretation "abstract pair" 'pair f g = (pair f g). + +axiom fst_pair: ∀a,b. fst 〈a,b〉 = a. +axiom snd_pair: ∀a,b. snd 〈a,b〉 = b. +axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉. + +axiom le_fst : ∀p. fst p ≤ p. +axiom le_snd : ∀p. snd p ≤ p. +axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉. + +(************************************* U **************************************) +axiom U: nat → nat →nat → option nat. + +axiom monotonic_U: ∀i,x,n,m,y.n ≤m → + U i x n = Some ? y → U i x m = Some ? y. + +lemma unique_U: ∀i,x,n,m,yn,ym. + U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. +#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. + +notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. + +lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n. +#i #x #n normalize cases (U i x n) + [%2 % * #y #H destruct|#y %1 %{y} //] +qed. + +lemma monotonic_terminate: ∀i,x,n,m. + n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m. +#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) // +qed. + +definition termb ≝ λi,x,t. + match U i x t with [None ⇒ false |Some y ⇒ true]. + +lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t. +#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] +qed. + +lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true. +#i #x #t * #y #H normalize >H // +qed. + +definition out ≝ λi,x,r. + match U i x r with [ None ⇒ 0 | Some z ⇒ z]. + +definition bool_to_nat: bool → nat ≝ + λb. match b with [true ⇒ 1 | false ⇒ 0]. + +coercion bool_to_nat. + +definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉. + +lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y. +#i #x #r #y % normalize + [cases (U i x r) normalize + [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H] + #H1 destruct + |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1] + #H1 // + ] + |#H >H //] +qed. + +lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?. +#i #x #r % normalize + [cases (U i x r) normalize // + #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1] + #H1 destruct + |#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. +======= +include "reverse_complexity/bigops_compl.ma". +>>>>>>> reverse_complexity lib restored:matita/matita/lib/reverse_complexity/speedup.ma +*) + +(********************************* the speedup ********************************) + +definition min_input ≝ λh,i,x. μ_{y ∈ [S i,x] } (termb i y (h (S i) y)). + +lemma min_input_def : ∀h,i,x. + min_input h i x = min (x -i) (S i) (λy.termb i y (h (S i) y)). +// qed. + +lemma min_input_i: ∀h,i,x. x ≤ i → min_input h i x = S i. +#h #i #x #lexi >min_input_def +cut (x - i = 0) [@sym_eq /2 by eq_minus_O/] #Hcut // +qed. + +lemma min_input_to_terminate: ∀h,i,x. + min_input h i x = x → {i ⊙ x} ↓ (h (S i) x). +#h #i #x #Hminx +cases (decidable_le (S i) x) #Hix + [cases (true_or_false (termb i x (h (S i) x))) #Hcase + [@termb_true_to_term // + |min_input_def in Hminx; #Hminx >Hminx in ⊢ (%→?); + min_input_i in Hminx; + [#eqix >eqix in Hix; * /2/ | @le_S_S_to_le @not_le_to_lt //] + ] +qed. + +lemma min_input_to_lt: ∀h,i,x. + min_input h i x = x → i < x. +#h #i #x #Hminx cases (decidable_le (S i) x) // +#ltxi @False_ind >min_input_i in Hminx; + [#eqix >eqix in ltxi; * /2/ | @le_S_S_to_le @not_le_to_lt //] +qed. + +lemma le_to_min_input: ∀h,i,x,x1. x ≤ x1 → + min_input h i x = x → min_input h i x1 = x. +#h #i #x #x1 #lex #Hminx @(min_exists … (le_S_S … lex)) + [@(fmin_true … (sym_eq … Hminx)) // + |@(min_input_to_lt … Hminx) + |#j #H1 g_def cut (x-u = 0) [/2 by minus_le_minus_minus_comm/] +#eq0 >eq0 normalize // qed. + +lemma g_lt : ∀h,i,x. min_input h i x = x → + out i x (h (S i) x) < g h 0 x. +#h #i #x #H @le_S_S @(le_MaxI … i) /2 by min_input_to_lt/ +qed. + +lemma max_neq0 : ∀a,b. max a b ≠ 0 → a ≠ 0 ∨ b ≠ 0. +#a #b whd in match (max a b); cases (true_or_false (leb a b)) #Hcase >Hcase + [#H %2 @H | #H %1 @H] +qed. + +definition almost_equal ≝ λf,g:nat → nat. ¬ ∀nu.∃x. nu < x ∧ f x ≠ g x. +interpretation "almost equal" 'napart f g = (almost_equal f g). + +lemma eventually_cancelled: ∀h,u.¬∀nu.∃x. nu < x ∧ + max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) ≠ 0. +#h #u elim u + [normalize % #H cases (H u) #x * #_ * #H1 @H1 // + |#u0 @not_to_not #Hind #nu cases (Hind nu) #x * #ltx + cases (true_or_false (eqb (min_input h (u0+O) x) x)) #Hcase + [>bigop_Strue [2:@Hcase] #Hmax cases (max_neq0 … Hmax) -Hmax + [2: #H %{x} % // bigop_Sfalse + [#H %{x1} % [@transitive_lt //| (le_to_min_input … (eqb_true_to_eq … Hcase)) + [@lt_to_not_eq @ltx1 | @lt_to_le @ltx1] + ] + |>bigop_Sfalse [2:@Hcase] #H %{x} % // (bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA) + [>H // |@lt_to_le @(le_to_lt_to_lt …ltx) /2 by le_maxr/ |//] +qed. + +(******************************** Condition 2 *********************************) + +lemma exists_to_exists_min: ∀h,i. (∃x. i < x ∧ {i ⊙ x} ↓ h (S i) x) → ∃y. min_input h i y = y. +#h #i * #x * #ltix #Hx %{(min_input h i x)} @min_spec_to_min @found // + [@(f_min_true (λy:ℕ.termb i y (h (S i) y))) %{x} % [% // | @term_to_termb_true //] + |#y #leiy #lty @(lt_min_to_false ????? lty) // + ] +qed. + +lemma condition_2: ∀h,i. code_for (total (g h 0)) i → ¬∃x. itermy >Hr +#H @(absurd ? H) @le_to_not_lt @le_n +qed. + +(**************************** complexity of g *********************************) + +definition unary_g ≝ λh.λux. g h (fst ux) (snd ux). +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 (auxg h) → CF s (unary_g h). +#h #s #H1 @(CF_compS ? (auxg h) H1) +qed. + +definition aux1g ≝ + λh,ux. max_{i ∈[fst ux,snd ux[ | (λp. eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) 〈i,ux〉} + ((λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) 〈i,ux〉). + +lemma eq_aux : ∀h,x.aux1g h x = auxg h x. +#h #x @same_bigop + [#n #_ >fst_pair >snd_pair // |#n #_ #_ >fst_pair >snd_pair //] +qed. + +lemma compl_g2 : ∀h,s1,s2,s. + CF s1 + (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) → + CF s2 + (λp:ℕ.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) → + O s (λx.MSC x + (snd x - fst x)*Max_{i ∈[fst x ,snd x[ }(s1 〈i,x〉+s2 〈i,x〉)) → + CF s (auxg h). +#h #s1 #s2 #s #Hs1 #Hs2 #HO @(ext_CF (aux1g h)) + [#n whd in ⊢ (??%%); @eq_aux] +@(CF_max2 … 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 (λ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 … (le_to_O … (MSC_in … H))) +@O_plus // %{1} %{0} #n #_ >commutative_times min_input_def whd in ⊢ (??%?); >minus_S_S @min_f_g #i #_ #_ +whd in ⊢ (??%%); >fst_pair >snd_pair // +qed. + +definition termb_aux ≝ λh. + termb_unary ∘ λp.〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉. + +lemma compl_g4 : ∀h,s1,s. (∀x. 0 < s1 x) → + (CF s1 + (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)))) → + (O s (λx.MSC x + ((snd(snd x) - (fst x)))*Max_{i ∈[S(fst x) ,S(snd (snd x))[ }(s1 〈i,x〉))) → + CF s (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #s1 #s #pos_s1 #Hs1 #HO @(ext_CF (min_input_aux h)) + [#n whd in ⊢ (??%%); @min_input_eq] +@(CF_mu4 … MSC MSC … pos_s1 … Hs1) + [@CF_compS @CF_fst + |@CF_comp_snd @CF_snd + |@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //] +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〉)). +#s1 #Hs1 %{1} %{0} #n #_ >commutative_times minus_S_S //] +qed. + +lemma coroll2: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) → +O (λx.(snd x - fst x)*s1 〈fst x,x〉) (λx.∑_{i ∈[fst x,snd x[ }(s1 〈i,x〉)). +#s1 #Hs1 %{1} %{0} #n #_ >commutative_times fst_pair >snd_pair >fst_pair >snd_pair // ] +@(CF_comp … (λx.MSC 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 + [@O_plus + [@(O_trans … (λx.MSC (fst x) + MSC (max (fst (snd x)) (snd (snd x))))) + [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx + >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) + >distributive_times_plus @le_plus [//] + cases (surj_pair b) #c * #d #eqb >eqb + >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) + whd in ⊢ (??%); @le_plus + [@monotonic_MSC @(le_maxl … (le_n …)) + |>commutative_times fst_pair >snd_pair // qed. + +lemma le_big : ∀x. x ≤ big x. +#x cases (surj_pair x) #a * #b #eqx >eqx @le_pair >fst_pair >snd_pair +[@(le_maxl … (le_n …)) | @(le_maxr … (le_n …))] +qed. + +definition faux2 ≝ λh. + (λx.MSC x + (snd (snd x)-fst x)* + (λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉). + +lemma compl_g7: ∀h. + constructible (λx. h (fst x) (snd x)) → + (∀n. monotonic ? le (h n)) → + CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉) + (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #hcostr #hmono @(monotonic_CF … (faux2 h)) + [#n normalize >fst_pair >snd_pair //] +@compl_g5 [#n @sU_pos | 3:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair +>fst_pair >snd_pair @monotonic_sU // @hmono @lexy +qed. + +lemma compl_g71: ∀h. + 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))〉〉) + (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #hcostr #hmono @(monotonic_CF … (compl_g7 h hcostr hmono)) #x +@le_plus [@monotonic_MSC //] +cases (decidable_le (fst x) (snd(snd x))) + [#Hle @le_times // @monotonic_sU + |#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt] + ] +qed. + +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 (MSC x)) >commutative_times @le_times + [// | @monotonic_MSC // ]] +@(O_trans … (coroll3 ??)) + [#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 [//|@(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 + [/2 by monotonic_le_minus_r/ + |@monotonic_sU // @hantimono [@le_S_S // |@ltb] + ] + |@monotonic_sU // @hantimono [@le_S_S // |@ltb] + ] + |@le_to_O #n >fst_pair >snd_pair + cut (max (fst n) n = n) [normalize >le_to_leb_true //] #Hmax >Hmax + >associative_plus >distributive_times_plus + @le_plus [@le_times [@le_S // |>big_def >Hmax //] |//] + ] +qed. + +definition c ≝ λx.(S (snd x-fst x))*MSC 〈x,x〉. + +definition sg ≝ λh,x. + let a ≝ fst x in + let b ≝ snd x in + c x + (b-a)*(S(b-a))*sU 〈x,〈snd x,h (S a) b〉〉. + +lemma sg_def1 : ∀h,a,b. + sg h 〈a,b〉 = c 〈a,b〉 + + (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉. +#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair // +qed. + +lemma sg_def : ∀h,a,b. + sg h 〈a,b〉 = + S (b-a)*MSC 〈〈a,b〉,〈a,b〉〉 + (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉. +#h #a #b normalize >fst_pair >snd_pair // +qed. + +lemma compl_g11 : ∀h. + 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) (unary_g h). +#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham) +qed. + +(**************************** closing the argument ****************************) + +let rec h_of_aux (r:nat →nat) (c,d,b:nat) on d : nat ≝ + match d with + [ O ⇒ c + | S d1 ⇒ (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + + d*(S d)*sU 〈〈b-d,b〉,〈b,r (h_of_aux r c d1 b)〉〉]. + +lemma h_of_aux_O: ∀r,c,b. + h_of_aux r c O b = c. +// qed. + +lemma h_of_aux_S : ∀r,c,d,b. + h_of_aux r c (S d) b = + (S (S d))*(MSC 〈〈b-(S d),b〉,〈b-(S d),b〉〉) + + (S d)*(S (S d))*sU 〈〈b-(S d),b〉,〈b,r(h_of_aux r c d b)〉〉. +// qed. + +definition h_of ≝ λr,p. + let m ≝ max (fst p) (snd p) in + h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (snd p - fst p) (snd p). + +lemma h_of_O: ∀r,a,b. b ≤ a → + h_of r 〈a,b〉 = let m ≝ max a b in MSC 〈〈m,m〉,〈m,m〉〉. +#r #a #b #Hle normalize >fst_pair >snd_pair >(minus_to_0 … Hle) // +qed. + +lemma h_of_def: ∀r,a,b.h_of r 〈a,b〉 = + let m ≝ max a b in + h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (b - a) b. +#r #a #b normalize >fst_pair >snd_pair // +qed. + +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. +#r #Hr #monor #d #d1 lapply d -d elim d1 + [#d #c #c1 #b #b1 #Hc #Hd @(le_n_O_elim ? Hd) #leb + >h_of_aux_O >h_of_aux_O // + |#m #Hind #d #c #c1 #b #b1 #lec #led #leb cases (le_to_or_lt_eq … led) + [#ltd @(transitive_le … (Hind … lec ? leb)) [@le_S_S_to_le @ltd] + >h_of_aux_S @(transitive_le ???? (le_plus_n …)) + >(times_n_1 (h_of_aux r c1 m b1)) in ⊢ (?%?); + >commutative_times @le_times [//|@(transitive_le … (Hr ?)) @sU_le] + |#Hd >Hd >h_of_aux_S >h_of_aux_S + cut (b-S m ≤ b1 - S m) [/2 by monotonic_le_minus_l/] #Hb1 + @le_plus [@le_times //] + [@monotonic_MSC @le_pair @le_pair // + |@le_times [//] @monotonic_sU + [@le_pair // |// |@monor @Hind //] + ] + ] + ] +qed. + +lemma mono_h_of2: ∀r.(∀x. x ≤ r x) → monotonic ? le r → + ∀i,b,b1. b ≤ b1 → h_of r 〈i,b〉 ≤ h_of r 〈i,b1〉. +#r #Hr #Hmono #i #a #b #leab >h_of_def >h_of_def +cut (max i a ≤ max i b) + [@to_max + [@(le_maxl … (le_n …))|@(transitive_le … leab) @(le_maxr … (le_n …))]] +#Hmax @(mono_h_of_aux r Hr Hmono) + [@monotonic_MSC @le_pair @le_pair @Hmax |/2 by monotonic_le_minus_l/ |@leab] +qed. + +axiom h_of_constr : ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + constructible (h_of r). + +lemma speed_compl: ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))). +#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …)) + [#x cases (surj_pair x) #a * #b #eqx >eqx + >sg_def cases (decidable_le b a) + [#leba >(minus_to_0 … leba) normalize in ⊢ (?%?); + h_of_def + cut (max a b = a) + [normalize cases (le_to_or_lt_eq … leba) + [#ltba >(lt_to_leb_false … ltba) % + |#eqba (le_to_leb_true … (le_n ?)) % ]] + #Hmax >Hmax normalize >(minus_to_0 … leba) normalize + @monotonic_MSC @le_pair @le_pair // + |#ltab >h_of_def >h_of_def + cut (max a b = b) + [normalize >(le_to_leb_true … ) [%] @lt_to_le @not_le_to_lt @ltab] + #Hmax >Hmax + 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) + [%{(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 + [@lt_to_le @not_le_to_lt // | //]] #eqd2 >eqd2 + normalize // + ] + |#n #a #b #leab #lebn >h_of_def >h_of_def + cut (max a n = n) + [normalize >le_to_leb_true [%|@(transitive_le … leab lebn)]] #Hmaxa + cut (max b n = n) + [normalize >(le_to_leb_true … lebn) %] #Hmaxb + >Hmaxa >Hmaxb @Hmono @(mono_h_of_aux r … Hr Hmono) // /2 by monotonic_le_minus_r/ + |#n #a #b #leab @Hmono @(mono_h_of2 … Hr Hmono … leab) + |@(constr_comp … Hconstr Hr) @(ext_constr (h_of r)) + [#x cases (surj_pair x) #a * #b #eqx >eqx >fst_pair >snd_pair //] + @(h_of_constr r Hr Hmono Hconstr) + ] +qed. + +lemma speed_compl_i: ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∀i. CF (λx.h_of r 〈i,x〉) (λx.g (λi,x. r(h_of r 〈i,x〉)) i x). +#r #Hr #Hmono #Hconstr #i +@(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 Hconstr)) #n // +qed. + +(**************************** the speedup theorem *****************************) +theorem pseudo_speedup: + ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∃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 #Hconstr +(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #i * +#Hcodei #HCi +(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} +(* sg is (λx.h_of r 〈i,x〉) *) +%{(λx. h_of r 〈S i,x〉)} +lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg +%[%[@condition_1 |@Hg] + |cases Hg #H1 * #j * #Hcodej #HCj + lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) + cases HCi #m * #a #Ha %{m} %{(max (S i) a)} #n #ltin @lt_to_le @not_le_to_lt + @(not_to_not … Hcond2) -Hcond2 #Hlesf %{n} % + [@(transitive_le … ltin) @(le_maxl … (le_n …))] + cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] + #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) // + ] +qed. + +theorem pseudo_speedup': + ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∃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 #Hconstr +(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #i * +#Hcodei #HCi +(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} +(* sg is (λx.h_of r 〈S i,x〉) *) +%{(λx. h_of r 〈S i,x〉)} +lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg +%[%[@condition_1 |@Hg] + |cases Hg #H1 * #j * #Hcodej #HCj + lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) + cases HCi #m * #a #Ha + %{m} %{(max (S i) a)} #n #ltin @not_le_to_lt @(not_to_not … Hcond2) -Hcond2 + #Hlesf %{n} % [@(transitive_le … ltin) @(le_maxl … (le_n …))] + cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] + #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) + @Hmono @(mono_h_of2 … Hr Hmono … ltin) + ] +qed. + diff --git a/matita/matita/lib/reverse_complexity/toolkit.ma b/matita/matita/lib/reverse_complexity/toolkit.ma new file mode 100644 index 000000000..665df181e --- /dev/null +++ b/matita/matita/lib/reverse_complexity/toolkit.ma @@ -0,0 +1,1623 @@ +include "basics/types.ma". +include "arithmetics/minimization.ma". +include "arithmetics/bigops.ma". +include "arithmetics/sigma_pi.ma". +include "arithmetics/bounded_quantifiers.ma". +include "reverse_complexity/big_O.ma". + +(************************* notation for minimization **************************) +notation "μ_{ ident i < n } p" + with precedence 80 for @{min $n 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ≤ n } p" + with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b[ } p" + with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}. + +notation "μ_{ ident i ∈ [a,b] } p" + with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}. + +(************************************ 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 le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true → + f a ≤ Max_{i ∈ [m,n[ | p i}(f i). +#f #p #n #m #a #lema #ltan #pa +>(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m)) + [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. + +(********************************** pairing ***********************************) +axiom pair: nat → nat → nat. +axiom fst : nat → nat. +axiom snd : nat → nat. + +interpretation "abstract pair" 'pair f g = (pair f g). + +axiom fst_pair: ∀a,b. fst 〈a,b〉 = a. +axiom snd_pair: ∀a,b. snd 〈a,b〉 = b. +axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉. + +axiom le_fst : ∀p. fst p ≤ p. +axiom le_snd : ∀p. snd p ≤ p. +axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉. + +lemma expand_pair: ∀x. x = 〈fst x, snd x〉. +#x lapply (surj_pair x) * #a * #b #Hx >Hx >fst_pair >snd_pair // +qed. + +(************************************* U **************************************) +axiom U: nat → nat →nat → option nat. + +axiom monotonic_U: ∀i,x,n,m,y.n ≤m → + U i x n = Some ? y → U i x m = Some ? y. + +lemma unique_U: ∀i,x,n,m,yn,ym. + U i x n = Some ? yn → U i x m = Some ? ym → yn = ym. +#i #x #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 ≝ λi,x,r. ∃y. U i x r = Some ? y. + +notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}. + +lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n. +#i #x #n normalize cases (U i x n) + [%2 % * #y #H destruct|#y %1 %{y} //] +qed. + +lemma monotonic_terminate: ∀i,x,n,m. + n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m. +#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) // +qed. + +definition termb ≝ λi,x,t. + match U i x t with [None ⇒ false |Some y ⇒ true]. + +lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t. +#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //] +qed. + +lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true. +#i #x #t * #y #H normalize >H // +qed. + +definition out ≝ λi,x,r. + match U i x r with [ None ⇒ 0 | Some z ⇒ z]. + +definition bool_to_nat: bool → nat ≝ + λb. match b with [true ⇒ 1 | false ⇒ 0]. + +coercion bool_to_nat. + +definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉. + +lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y. +#i #x #r #y % normalize + [cases (U i x r) normalize + [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H] + #H1 destruct + |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1] + #H1 // + ] + |#H >H //] +qed. + +lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?. +#i #x #r % normalize + [cases (U i x r) normalize // + #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1] + #H1 destruct + |#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. + + +definition total ≝ λf.λx:nat. Some nat (f x). + + +(********************************* complexity *********************************) + +(* We assume operations have a minimal structural complexity MSC. +For instance, for time complexity, MSC is equal to the size of input. +For space complexity, MSC is typically 0, since we only measure the +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) *) + +definition C ≝ λs,i.∃c.∃a.∀x.a ≤ x → ∃y. + U i x (c*(s x)) = Some ? 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 | //] +qed. + +lemma monotonic_CF: ∀s1,s2,f.(∀x. s1 x ≤ s2 x) → CF s1 f → CF s2 f. +#s1 #s2 #f #H * #HO * #i * #Hcode #Hs1 % + [cases HO #c * #a -HO #HO %{c} %{a} #n #lean @(transitive_le … (HO n lean)) + @le_times // + |%{i} % [//] cases Hs1 #c * #a -Hs1 #Hs1 %{c} %{a} #n #lean + cases(Hs1 n lean) #y #Hy %{y} @(monotonic_U …Hy) @le_times // + ] +qed. + +lemma O_to_CF: ∀s1,s2,f.O s2 s1 → CF s1 f → CF s2 f. +#s1 #s2 #f #H * #HO * #i * #Hcode #Hs1 % + [@(O_trans … H) // + |%{i} % [//] cases Hs1 #c * #a -Hs1 #Hs1 + cases H #c1 * #a1 #Ha1 %{(c*c1)} %{(a+a1)} #n #lean + cases(Hs1 n ?) [2:@(transitive_le … lean) //] #y #Hy %{y} @(monotonic_U …Hy) + >associative_times @le_times // @Ha1 @(transitive_le … lean) // + ] +qed. + +lemma timesc_CF: ∀s,f,c.CF (λx.c*s x) f → CF s f. +#s #f #c @O_to_CF @O_times_c +qed. + +(********************************* composition ********************************) +axiom CF_comp: ∀f,g,sf,sg,sh. CF sg g → CF sf f → + O sh (λx. sg x + sf (g x)) → CF sh (f ∘ g). + +lemma CF_comp_ext: ∀f,g,h,sh,sf,sg. CF sg g → CF sf f → + (∀x.f(g x) = h x) → O sh (λx. sg x + sf (g x)) → CF sh h. +#f #g #h #sh #sf #sg #Hg #Hf #Heq #H @(ext_CF (f ∘ g)) + [#n normalize @Heq | @(CF_comp … H) //] +qed. + +(* primitve recursion *) + +(* no arguments *) + +let rec prim_rec0 (k:nat) (h:nat →nat) n on n ≝ + match n with + [ O ⇒ k + | S a ⇒ h 〈a, prim_rec0 k h a〉]. + +lemma prim_rec0_S: ∀k,h,n. + prim_rec0 k h (S n) = h 〈n, prim_rec0 k h n〉. +// qed. + +let rec prim_rec0_compl (k,sk:nat) (h,sh:nat →nat) n on n ≝ + match n with + [ O ⇒ sk + | S a ⇒ prim_rec0_compl k sk h sh a + sh (prim_rec0 k h a)]. + +axiom CF_prim_rec0_gen: ∀k,h,sk,sh,sh1,sf. CF sh h → + O sh1 (λx.snd x + sh 〈fst x,prim_rec0 k h (fst x)〉) → + O sf (prim_rec0 sk sh1) → + CF sf (prim_rec0 k h). + +lemma CF_prim_rec0: ∀k,h,sk,sh,sf. CF sh h → + O sf (prim_rec0 sk (λx. snd x + sh 〈fst x,prim_rec0 k h (fst x)〉)) + → CF sf (prim_rec0 k h). +#k #h #sk #sh #sf #Hh #HO @(CF_prim_rec0_gen … Hh … HO) @O_refl +qed. + +(* with argument(s) m *) +let rec prim_rec (k,h:nat →nat) n m on n ≝ + match n with + [ O ⇒ k m + | S a ⇒ h 〈a,〈prim_rec k h a m, m〉〉]. + +lemma prim_rec_S: ∀k,h,n,m. + prim_rec k h (S n) m = h 〈n,〈prim_rec k h n m, m〉〉. +// qed. + +lemma prim_rec_le: ∀k1,k2,h1,h2. (∀x.k1 x ≤ k2 x) → +(∀x,y.x ≤y → h1 x ≤ h2 y) → + ∀x,m. prim_rec k1 h1 x m ≤ prim_rec k2 h2 x m. +#k1 #k2 #h1 #h2 #lek #leh #x #m elim x // +#n #Hind normalize @leh @le_pair // @le_pair // +qed. + +definition unary_pr ≝ λk,h,x. prim_rec k h (fst x) (snd x). + +lemma prim_rec_unary: ∀k,h,a,b. + prim_rec k h a b = unary_pr k h 〈a,b〉. +#k #h #a #b normalize >fst_pair >snd_pair // +qed. + + +let rec prim_rec_compl (k,h,sk,sh:nat →nat) n m on n ≝ + match n with + [ O ⇒ sk m + | S a ⇒ prim_rec_compl k h sk sh a m + sh (prim_rec k h a m)]. + +axiom CF_prim_rec_gen: ∀k,h,sk,sh,sh1. CF sk k → CF sh h → + O sh1 (λx. fst (snd x) + sh 〈fst x,〈unary_pr k h 〈fst x,snd (snd x)〉,snd (snd x)〉〉) → + CF (unary_pr sk sh1) (unary_pr k h). + +lemma CF_prim_rec: ∀k,h,sk,sh,sf. CF sk k → CF sh h → + O sf (unary_pr sk (λx. fst (snd x) + sh 〈fst x,〈unary_pr k h 〈fst x,snd (snd x)〉,snd (snd x)〉〉)) + → CF sf (unary_pr k h). +#k #h #sk #sh #sf #Hk #Hh #Osf @(O_to_CF … Osf) @(CF_prim_rec_gen … Hk Hh) @O_refl +qed. + +(**************************** primitive operations*****************************) + +definition id ≝ λx:nat.x. + +axiom CF_id: CF MSC id. +axiom CF_const: ∀k.CF MSC (λx.k). +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_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 ***********************************) + +axiom CF_eqb: ∀h,f,g. + CF h f → CF h g → CF h (λx.eqb (f x) (g x)). + +(*********************************** maximum **********************************) + +alias symbol "pair" (instance 13) = "abstract pair". +alias symbol "pair" (instance 12) = "abstract pair". +alias symbol "and" (instance 11) = "boolean and". +alias symbol "plus" (instance 8) = "natural plus". +alias symbol "pair" (instance 7) = "abstract pair". +alias symbol "plus" (instance 6) = "natural plus". +alias symbol "pair" (instance 5) = "abstract pair". +alias id "max" = "cic:/matita/arithmetics/nat/max#def:2". +alias symbol "minus" (instance 3) = "natural minus". +lemma max_gen: ∀a,b,p,f,x. a ≤b → + max_{i ∈[a,b[ | p 〈i,x〉 }(f 〈i,x〉) = max_{i < b | leb a i ∧ p 〈i,x〉 }(f 〈i,x〉). +#a #b #p #f #x @(bigop_I_gen ????? MaxA) +qed. + +lemma max_prim_rec_base: ∀a,b,p,f,x. a ≤b → + max_{i < b| p 〈i,x〉 }(f 〈i,x〉) = + prim_rec (λi.0) + (λi.if p 〈fst i,x〉 then max (f 〈fst i,snd (snd i)〉) (fst (snd i)) else fst (snd i)) b x. +#a #b #p #f #x #leab >max_gen // elim b + [normalize // + |#i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair + cases (true_or_false (p 〈i,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma max_prim_rec: ∀a,b,p,f,x. a ≤b → + max_{i ∈[a,b[ | p 〈i,x〉 }(f 〈i,x〉) = + prim_rec (λi.0) + (λi.if leb a (fst i) ∧ p 〈fst i,x〉 then max (f 〈fst i,snd (snd i)〉) (fst (snd i)) else fst (snd i)) b x. +#a #b #p #f #x #leab >max_gen // elim b + [normalize // + |#i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair + cases (true_or_false (leb a i ∧ p 〈i,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma max_prim_rec1: ∀a,b,p,f,x. + max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) = + prim_rec (λi.0) + (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉 + then max (f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉) (fst (snd i)) + else fst (snd i)) (b x-a x) 〈a x ,x〉. +#a #b #p #f #x elim (b x-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma sum_prim_rec1: ∀a,b,p,f,x. + ∑_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) = + prim_rec (λi.0) + (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉 + then plus (f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉) (fst (snd i)) + else fst (snd i)) (b x-a x) 〈a x ,x〉. +#a #b #p #f #x elim (b x-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma bigop_prim_rec: ∀a,b,c,p,f,x. + bigop (b x-a x) (λi:ℕ.p 〈i+a x,x〉) ? (c 〈a x,x〉) plus (λi:ℕ.f 〈i+a x,x〉) = + prim_rec c + (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉 + then plus (f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉) (fst (snd i)) + else fst (snd i)) (b x-a x) 〈a x ,x〉. +#a #b #c #p #f #x normalize elim (b x-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma bigop_prim_rec_dec: ∀a,b,c,p,f,x. + bigop (b x-a x) (λi:ℕ.p 〈b x -i,x〉) ? (c 〈b x,x〉) plus (λi:ℕ.f 〈b x-i,x〉) = + prim_rec c + (λi.if p 〈fst (snd (snd i)) - fst i,snd (snd (snd i))〉 + then plus (f 〈fst (snd (snd i)) - fst i,snd (snd (snd i))〉) (fst (snd i)) + else fst (snd i)) (b x-a x) 〈b x ,x〉. +#a #b #c #p #f #x normalize elim (b x-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + cases (true_or_false (p 〈b x - i,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +lemma bigop_prim_rec_dec1: ∀a,b,c,p,f,x. + bigop (S(b x)-a x) (λi:ℕ.p 〈b x - i,x〉) ? (c 〈b x,x〉) plus (λi:ℕ.f 〈b x- i,x〉) = + prim_rec c + (λi.if p 〈fst (snd (snd i)) - (fst i),snd (snd (snd i))〉 + then plus (f 〈fst (snd (snd i)) - (fst i),snd (snd (snd i))〉) (fst (snd i)) + else fst (snd i)) (S(b x)-a x) 〈b x,x〉. +#a #b #c #p #f #x elim (S(b x)-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + cases (true_or_false (p 〈b x - i,x〉)) #Hcase >Hcase + [bigop_Strue // |bigop_Sfalse // ] + ] +qed. + +(* +lemma bigop_aux_1: ∀k,c,f. + bigop (S k) (λi:ℕ.true) ? c plus (λi:ℕ.f i) = + f 0 + bigop k (λi:ℕ.true) ? c plus (λi:ℕ.f (S i)). +#k #c #f elim k [normalize //] #i #Hind >bigop_Strue // + +lemma bigop_prim_rec_dec: ∀a,b,c,f,x. + bigop (b x-a x) (λi:ℕ.true) ? (c 〈b x,x〉) plus (λi:ℕ.f 〈i+a x,x〉) = + prim_rec c + (λi. plus (f 〈fst (snd (snd i)) - fst i,snd (snd (snd i))〉) (fst (snd i))) + (b x-a x) 〈b x ,x〉. +#a #b #c #f #x normalize elim (b x-a x) + [normalize // + |#i #Hind >prim_rec_S + >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair + bigop_Strue // |bigop_Sfalse // ] + ] +qed. *) + +lemma bigop_plus_c: ∀k,p,f,c. + c k + bigop k (λi.p i) ? 0 plus (λi.f i) = + bigop k (λi.p i) ? (c k) plus (λi.f i). +#k #p #f elim k [normalize //] +#i #Hind #c cases (true_or_false (p i)) #Hcase +[>bigop_Strue // >bigop_Strue // (commutative_plus ? (f i)) + >associative_plus @eq_f @Hind +|>bigop_Sfalse // >bigop_Sfalse // +] +qed. + +(* the argument is 〈b-a,〈a,x〉〉 *) + + +definition max_unary_pr ≝ λp,f.unary_pr (λx.0) + (λi. + let k ≝ fst i in + let r ≝ fst (snd i) in + let a ≝ fst (snd (snd i)) in + let x ≝ snd (snd (snd i)) in + if p 〈k + a,x〉 then max (f 〈k+a,x〉) r else r). + +lemma max_unary_pr1: ∀a,b,p,f,x. + max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) = + ((max_unary_pr p f) ∘ (λx.〈b x - a x,〈a x,x〉〉)) x. +#a #b #p #f #x normalize >fst_pair >snd_pair @max_prim_rec1 +qed. + +(* +lemma max_unary_pr: ∀a,b,p,f,x. + max_{i ∈[a,b[ | p 〈i,x〉 }(f 〈i,x〉) = + prim_rec (λi.0) + (λi.if p 〈fst i +a,x〉 then max (f 〈fst i +a ,snd (snd i)〉) (fst (snd i)) else fst (snd i)) (b-a) x. +*) + +(* +definition unary_compl ≝ λp,f,hf. + unary_pr MSC + (λx:ℕ + .fst (snd x) + +hf + 〈fst x, + 〈unary_pr (λx0:ℕ.O) + (λi:ℕ + .(let (k:ℕ) ≝fst i in  + let (r:ℕ) ≝fst (snd i) in  + let (a:ℕ) ≝fst (snd (snd i)) in  + let (x:ℕ) ≝snd (snd (snd i)) in  + if p 〈k+a,x〉 then max (f 〈k+a,x〉) r else r )) 〈fst x,snd (snd x)〉, + snd (snd x)〉〉). *) + +definition aux_compl ≝ λhp,hf.λi. + let k ≝ fst i in  + let r ≝ fst (snd i) in  + let a ≝ fst (snd (snd i)) in  + let x ≝ snd (snd (snd i)) in  + hp 〈k+a,x〉 + hf 〈k+a,x〉 + (* was MSC r*) MSC i . + +definition aux_compl1 ≝ λhp,hf.λi. + let k ≝ fst i in  + let r ≝ fst (snd i) in  + let a ≝ fst (snd (snd i)) in  + let x ≝ snd (snd (snd i)) in  + hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r. + +lemma aux_compl1_def: ∀k,r,m,hp,hf. + aux_compl1 hp hf 〈k,〈r,m〉〉 = + let a ≝ fst m in  + let x ≝ snd m in  + hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r. +#k #r #m #hp #hf normalize >fst_pair >snd_pair >snd_pair >fst_pair // +qed. + +lemma aux_compl1_def1: ∀k,r,a,x,hp,hf. + aux_compl1 hp hf 〈k,〈r,〈a,x〉〉〉 = hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r. +#k #r #a #x #hp #hf normalize >fst_pair >snd_pair >snd_pair >fst_pair +>fst_pair >snd_pair // +qed. + + +axiom Oaux_compl: ∀hp,hf. O (aux_compl1 hp hf) (aux_compl hp hf). + +(* +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 CF_if: ∀b:nat → bool. ∀f,g,s. CF s b → CF s f → CF s g → + CF s (λx. if b x then f x else g x). + +(* +lemma 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). +#b #f #g #sb #sf #sg #Hb #Hf #Hg @IF_CF_new + [@(monotonic_CF … Hb) @O_plus_l @O_plus_l // + |@(monotonic_CF … Hf) @O_plus_l @O_plus_r // + |@(monotonic_CF … Hg) @O_plus_r // + ] +qed. +*) + +axiom CF_le: ∀h,f,g. CF h f → CF h g → CF h (λx. leb (f x) (g x)). +axiom CF_max1: ∀h,f,g. CF h f → CF h g → CF h (λx. max (f x) (g x)). +axiom CF_plus: ∀h,f,g. CF h f → CF h g → CF h (λx. f x + g x). +axiom CF_minus: ∀h,f,g. CF h f → CF h g → CF h (λx. f x - g x). + +axiom MSC_prop: ∀f,h. CF h f → ∀x. MSC (f x) ≤ h x. + +lemma MSC_max: ∀f,h,x. CF h f → MSC (max_{i < x}(f i)) ≤ max_{i < x}(h i). +#f #h #x #hf elim x // #i #Hind >bigop_Strue [|//] >bigop_Strue [|//] +whd in match (max ??); +cases (true_or_false (leb (f i) (bigop i (λi0:ℕ.true) ? 0 max(λi0:ℕ.f i0)))) +#Hcase >Hcase + [@(transitive_le … Hind) @le_maxr // + |@(transitive_le … (MSC_prop … hf i)) @le_maxl // + ] +qed. + +lemma CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s. + CF ha a → CF hb b → CF hp p → CF hf f → + O s (λx.ha x + hb x + + (∑_{i ∈[a x ,b x[ } + (hp 〈i,x〉 + hf 〈i,x〉 + max_{i ∈ [a x, b x [ }(hf 〈i,x〉)))) → + CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). +#a #b #p #f #ha #hb #hp #hf #s #CFa #CFb #CFp #CFf #HO +@ext_CF1 [|#x @max_unary_pr1] +@(CF_comp ??? (λx.ha x + hb x)) + [|@CF_comp_pair + [@CF_minus [@(O_to_CF … CFb) @O_plus_r // |@(O_to_CF … CFa) @O_plus_l //] + |@CF_comp_pair + [@(O_to_CF … CFa) @O_plus_l // + | @(O_to_CF … CF_id) @O_plus_r @(proj1 … CFb) + ] + ] + |@(CF_prim_rec … MSC … (aux_compl1 hp hf)) + [@CF_const + |@(O_to_CF … (Oaux_compl … )) + @CF_if + [@(CF_comp p … MSC … CFp) + [@CF_comp_pair + [@CF_plus [@CF_fst| @CF_comp_fst @CF_comp_snd @CF_snd] + |@CF_comp_snd @CF_comp_snd @CF_snd] + |@le_to_O #x normalize >commutative_plus >associative_plus @le_plus // + ] + |@CF_max1 + [@(CF_comp f … MSC … CFf) + [@CF_comp_pair + [@CF_plus [@CF_fst| @CF_comp_fst @CF_comp_snd @CF_snd] + |@CF_comp_snd @CF_comp_snd @CF_snd] + |@le_to_O #x normalize >commutative_plus // + ] + |@CF_comp_fst @(monotonic_CF … CF_snd) normalize // + ] + |@CF_comp_fst @(monotonic_CF … CF_snd) normalize // + ] + |@O_refl + ] + |@(O_trans … HO) -HO + @(O_trans ? (λx:ℕ + .ha x+hb x + +bigop (b x-a x) (λi:ℕ.true) ? (MSC 〈a x,x〉) plus + (λi:ℕ + .(λi0:ℕ + .hp 〈i0,x〉+hf 〈i0,x〉 + +bigop (b x-a x) (λi1:ℕ.true) ? 0 max + (λi1:ℕ.(λi2:ℕ.hf 〈i2,x〉) (i1+a x))) (i+a x)))) + [ + @le_to_O #n @le_plus // whd in match (unary_pr ???); + >fst_pair >snd_pair >bigop_prim_rec elim (b n - a n) + [normalize // + |#x #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >aux_compl1_def1 + >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >fst_pair + >snd_pair normalize in ⊢ (??%); >commutative_plus @le_plus + [-Hind @le_plus // normalize >fst_pair >snd_pair + @(transitive_le ? (bigop x (λi1:ℕ.true) ? 0 (λn0:ℕ.λm:ℕ.if leb n0 m then m else n0 ) + (λi1:ℕ.hf 〈i1+a n,n〉))) + [elim x [normalize @MSC_le] + #x0 #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >snd_pair + >fst_pair >fst_pair cases (p 〈x0+a n,n〉) normalize + [cases (true_or_false (leb (f 〈x0+a n,n〉) + (prim_rec (λx00:ℕ.O) + (λi:ℕ + .if p 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉  + then if leb (f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉) + (fst (snd i))  + then fst (snd i)  + else f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉   + else fst (snd i) ) x0 〈a n,n〉))) #Hcase >Hcase normalize + [@(transitive_le … Hind) -Hind @(le_maxr … (le_n …)) + |@(transitive_le … (MSC_prop … CFf …)) @(le_maxl … (le_n …)) + ] + |@(transitive_le … Hind) -Hind @(le_maxr … (le_n …)) + ] + |@(le_maxr … (le_n …)) + ] + |@(transitive_le … Hind) -Hind + generalize in match (bigop x (λi:ℕ.true) ? 0 max + (λi1:ℕ.(λi2:ℕ.hf 〈i2,n〉) (i1+a n))); #c + generalize in match (hf 〈x+a n,n〉); #c1 + elim x [//] #x0 #Hind + >prim_rec_S >prim_rec_S normalize >fst_pair >fst_pair >snd_pair + >snd_pair >snd_pair >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair + >fst_pair @le_plus + [@le_plus // cases (true_or_false (leb c1 c)) #Hcase + >Hcase normalize // @lt_to_le @not_le_to_lt @(leb_false_to_not_le ?? Hcase) + |@Hind + ] + ] + ] + |@O_plus [@O_plus_l //] @le_to_O #x + Hind +cases (true_or_false (eqb a (a+i))) #Hcase >Hcase normalize // +<(eqb_true_to_eq … Hcase) >H // +qed. + +lemma my_minim_nfa : ∀a,f,x,k. f〈a,x〉 = false → +my_minim a f x (S k) = my_minim (S a) f x k. +#a #f #x #k #H elim k + [normalize H >eq_to_eqb_true // + |#i #Hind >my_minim_S >Hind >my_minim_S Hcase + [>(my_minim_fa … Hcase) // | >Hind @sym_eq @(my_minim_nfa … Hcase) ] +qed. + +(* cambiare qui: togliere S *) + + +definition minim_unary_pr ≝ λf.unary_pr (λx.S(fst x)) + (λi. + let k ≝ fst i in + let r ≝ fst (snd i) in + let b ≝ fst (snd (snd i)) in + let x ≝ snd (snd (snd i)) in + if f 〈b-k,x〉 then b-k else r). + +definition compl_minim_unary ≝ λhf:nat → nat.λi. + let k ≝ fst i in + let b ≝ fst (snd (snd i)) in + let x ≝ snd (snd (snd i)) in + hf 〈b-k,x〉 + MSC 〈k,〈S b,x〉〉. + +definition compl_minim_unary_aux ≝ λhf,i. + let k ≝ fst i in + let r ≝ fst (snd i) in + let b ≝ fst (snd (snd i)) in + let x ≝ snd (snd (snd i)) in + hf 〈b-k,x〉 + MSC i. + +lemma compl_minim_unary_aux_def : ∀hf,k,r,b,x. + compl_minim_unary_aux hf 〈k,〈r,〈b,x〉〉〉 = hf 〈b-k,x〉 + MSC 〈k,〈r,〈b,x〉〉〉. +#hf #k #r #b #x normalize >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair // +qed. + +lemma compl_minim_unary_def : ∀hf,k,r,b,x. + compl_minim_unary hf 〈k,〈r,〈b,x〉〉〉 = hf 〈b-k,x〉 + MSC 〈k,〈S b,x〉〉. +#hf #k #r #b #x normalize >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair // +qed. + +lemma compl_minim_unary_aux_def2 : ∀hf,k,r,x. + compl_minim_unary_aux hf 〈k,〈r,x〉〉 = hf 〈fst x-k,snd x〉 + MSC 〈k,〈r,x〉〉. +#hf #k #r #x normalize >snd_pair >snd_pair >fst_pair // +qed. + +lemma compl_minim_unary_def2 : ∀hf,k,r,x. + compl_minim_unary hf 〈k,〈r,x〉〉 = hf 〈fst x-k,snd x〉 + MSC 〈k,〈S(fst x),snd x〉〉. +#hf #k #r #x normalize >snd_pair >snd_pair >fst_pair // +qed. + +lemma min_aux: ∀a,f,x,k. min (S k) (a x) (λi:ℕ.f 〈i,x〉) = + ((minim_unary_pr f) ∘ (λx.〈S k,〈a x + k,x〉〉)) x. +#a #f #x #k whd in ⊢ (???%); >fst_pair >snd_pair +lapply a -a (* @max_prim_rec1 *) +elim k + [normalize #a >fst_pair >snd_pair >fst_pair >snd_pair >snd_pair >fst_pair + fst_pair // + |#i #Hind #a normalize in ⊢ (??%?); >prim_rec_S >fst_pair >snd_pair + >fst_pair >snd_pair >snd_pair >fst_pair minus_S_S <(minus_plus_m_m (a x) i) // +qed. + +lemma minim_unary_pr1: ∀a,b,f,x. + μ_{i ∈[a x,b x]}(f 〈i,x〉) = + if leb (a x) (b x) then ((minim_unary_pr f) ∘ (λx.〈S (b x) - a x,〈b x,x〉〉)) x + else (a x). +#a #b #f #x cases (true_or_false (leb (a x) (b x))) #Hcase >Hcase + [cut (b x = a x + (b x - a x)) + [>commutative_plus minus_Sn_m [|@leb_true_to_le //] + >min_aux whd in ⊢ (??%?); eq_minus_O // @not_le_to_lt @leb_false_to_not_le // + ] +qed. + +axiom sum_inv: ∀a,b,f.∑_{i ∈ [a,S b[ }(f i) = ∑_{i ∈ [a,S b[ }(f (a + b - i)). + +(* +#a #b #f @(bigop_iso … plusAC) whd %{(λi.b - a - i)} %{(λi.b - a -i)} % + [%[#i #lti #_ normalize @eq_f >commutative_plus commutative_plus // + ] + |@(O_to_CF … MSC) + [@le_to_O #x normalize // + |@CF_minus + [@CF_comp_fst @CF_comp_snd @CF_snd |@CF_fst] + ] + |@CF_comp_fst @(monotonic_CF … CF_snd) normalize // + ] + |@O_plus + [@O_plus_l @O_refl + |@O_plus_r @O_ext2 [||#x >compl_minim_unary_aux_def2 @refl] + @O_trans [||@le_to_O #x >compl_minim_unary_def2 @le_n] + @O_plus [@O_plus_l //] + @O_plus_r + @O_trans [|@le_to_O #x @MSC_pair] @O_plus + [@le_to_O #x @monotonic_MSC @(transitive_le ???? (le_fst …)) + >fst_pair @le_n] + @O_trans [|@le_to_O #x @MSC_pair] @O_plus + [@le_to_O #x @monotonic_MSC @(transitive_le ???? (le_snd …)) + >snd_pair @(transitive_le ???? (le_fst …)) >fst_pair + normalize >snd_pair >fst_pair lapply (surj_pair x) + * #x1 * #x2 #Hx >Hx >fst_pair >snd_pair elim x1 // + #i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair + cases (f ?) [@le_S // | //]] + @le_to_O #x @monotonic_MSC @(transitive_le ???? (le_snd …)) >snd_pair + >(expand_pair (snd (snd x))) in ⊢ (?%?); @le_pair // + ] + ] + |cut (O s (λx.ha x + hb x + + ∑_{i ∈[a x ,S(b x)[ }(hf 〈a x+b x-i,x〉 + MSC 〈b x -(a x+b x-i),〈S(b x),x〉〉))) + [@(O_ext2 … HO) #x @eq_f @sum_inv] -HO #HO + @(O_trans … HO) -HO + @(O_trans ? (λx:ℕ.ha x+hb x + +bigop (S(b x)-a x) (λi:ℕ.true) ? + (MSC 〈b x,x〉) plus (λi:ℕ.(λj.hf j + MSC 〈b x - fst j,〈S(b (snd j)),snd j〉〉) 〈b x- i,x〉))) + [@le_to_O #n @le_plus // whd in match (unary_pr ???); + >fst_pair >snd_pair >(bigop_prim_rec_dec1 a b ? (λi.true)) + (* it is crucial to recall that the index is bound by S(b x) *) + cut (S (b n) - a n ≤ S (b n)) [//] + elim (S(b n) - a n) + [normalize // + |#x #Hind #lex >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair + >compl_minim_unary_def >prim_rec_S >fst_pair >snd_pair >fst_pair + >snd_pair >fst_pair >snd_pair >fst_pair whd in ⊢ (??%); >commutative_plus + @le_plus [2:@Hind @le_S @le_S_S_to_le @lex] -Hind >snd_pair + >minus_minus_associative // @le_S_S_to_le // + ] + |@O_plus [@O_plus_l //] @O_ext2 [||#x @bigop_plus_c] + @O_plus + [@O_plus_l @O_trans [|@le_to_O #x @MSC_pair] + @O_plus [@O_plus_r @le_to_O @(MSC_prop … CFb)|@O_plus_r @(proj1 … CFb)] + |@O_plus_r @(O_ext2 … (O_refl …)) #x @same_bigop + [//|#i #H #_ @eq_f2 [@eq_f @eq_f2 //|>fst_pair @eq_f @eq_f2 //] + ] + ] + ] + ] + |@(O_to_CF … CFa) @(O_trans … HO) @O_plus_l @O_plus_l @O_refl + ] +qed. + +(* +lemma 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〉)). +#a #b #f #ha #hb #hf #s #CFa #CFb #CFf #HO +@ext_CF1 [|#x @minim_unary_pr1] +@(CF_comp ??? (λx.ha x + hb x)) + [|@CF_comp_pair + [@CF_minus [@CF_compS @(O_to_CF … CFb) @O_plus_r // |@(O_to_CF … CFa) @O_plus_l //] + |@CF_comp_pair + [@CF_max1 [@(O_to_CF … CFa) @O_plus_l // | @CF_compS @(O_to_CF … CFb) @O_plus_r //] + | @(O_to_CF … CF_id) @O_plus_r @(proj1 … CFb) + ] + ] + |@(CF_prim_rec … MSC … (compl_minim_unary_aux hf)) + [@CF_fst + |@CF_if + [@(CF_comp f … MSC … CFf) + [@CF_comp_pair + [@CF_minus [@CF_comp_fst @CF_comp_snd @CF_snd|@CF_compS @CF_fst] + |@CF_comp_snd @CF_comp_snd @CF_snd] + |@le_to_O #x normalize >commutative_plus // + ] + |@(O_to_CF … MSC) + [@le_to_O #x normalize // + |@CF_minus + [@CF_comp_fst @CF_comp_snd @CF_snd |@CF_compS @CF_fst] + ] + |@CF_comp_fst @(monotonic_CF … CF_snd) normalize // + ] + |@O_refl + ] + |@(O_trans … HO) -HO + @(O_trans ? (λx:ℕ + .ha x+hb x + +bigop (S(b x)-a x) (λi:ℕ.true) ? (MSC 〈a x,x〉) plus (λi:ℕ.hf 〈i+a x,x〉))) + [@le_to_O #n @le_plus // whd in match (unary_pr ???); + >fst_pair >snd_pair >(bigop_prim_rec ? (λn.S(b n)) ? (λi.true)) elim (S(b n) - a n) + [normalize @monotonic_MSC @le_pair + |#x #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >compl_minim_unary_def + >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >fst_pair + >snd_pair normalize in ⊢ (??%); >commutative_plus @le_plus + [-Hind @le_plus // normalize >fst_pair >snd_pair + @(transitive_le ? (bigop x (λi1:ℕ.true) ? 0 (λn0:ℕ.λm:ℕ.if leb n0 m then m else n0 ) + (λi1:ℕ.hf 〈i1+a n,n〉))) + [elim x [normalize @MSC_le] + #x0 #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >snd_pair + >fst_pair >fst_pair cases (p 〈x0+a n,n〉) normalize + [cases (true_or_false (leb (f 〈x0+a n,n〉) + (prim_rec (λx00:ℕ.O) + (λi:ℕ + .if p 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉  + then if leb (f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉) + (fst (snd i))  + then fst (snd i)  + else f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉   + else fst (snd i) ) x0 〈a n,n〉))) #Hcase >Hcase normalize + [@(transitive_le … Hind) -Hind @(le_maxr … (le_n …)) + |@(transitive_le … (MSC_prop … CFf …)) @(le_maxl … (le_n …)) + ] + |@(transitive_le … Hind) -Hind @(le_maxr … (le_n …)) + ] + |@(le_maxr … (le_n …)) + ] + + + @(O_trans ? (λx:ℕ + .ha x+hb x + +bigop (b x-a x) (λi:ℕ.true) ? (MSC 〈a x,x〉) plus + (λi:ℕ + .(λi0:ℕ + .hp 〈i0,x〉+hf 〈i0,x〉 + +bigop (b x-a x) (λi1:ℕ.true) ? 0 max + (λi1:ℕ.(λi2:ℕ.hf 〈i2,x〉) (i1+a x))) (i+a x)))) + [ + @le_to_O #n @le_plus // whd in match (unary_pr ???); + >fst_pair >snd_pair >bigop_prim_rec elim (b n - a n) + [normalize // + |#x #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >aux_compl1_def1 + >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >fst_pair + >snd_pair normalize in ⊢ (??%); >commutative_plus @le_plus + [-Hind @le_plus // normalize >fst_pair >snd_pair + @(transitive_le ? (bigop x (λi1:ℕ.true) ? 0 (λn0:ℕ.λm:ℕ.if leb n0 m then m else n0 ) + (λi1:ℕ.hf 〈i1+a n,n〉))) + [elim x [normalize @MSC_le] + #x0 #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >snd_pair + >fst_pair >fst_pair cases (p 〈x0+a n,n〉) normalize + [cases (true_or_false (leb (f 〈x0+a n,n〉) + (prim_rec (λx00:ℕ.O) + (λi:ℕ + .if p 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉  + then if leb (f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉) + (fst (snd i))  + then fst (snd i)  + else f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉   + else fst (snd i) ) x0 〈a n,n〉))) #Hcase >Hcase normalize + [@(transitive_le … Hind) -Hind @(le_maxr … (le_n …)) + |@(transitive_le … (MSC_prop … CFf …)) @(le_maxl … (le_n …)) + ] + |@(transitive_le … Hind) -Hind @(le_maxr … (le_n …)) + ] + |@(le_maxr … (le_n …)) + ] + |@(transitive_le … Hind) -Hind + generalize in match (bigop x (λi:ℕ.true) ? 0 max + (λi1:ℕ.(λi2:ℕ.hf 〈i2,n〉) (i1+a n))); #c + generalize in match (hf 〈x+a n,n〉); #c1 + elim x [//] #x0 #Hind + >prim_rec_S >prim_rec_S normalize >fst_pair >fst_pair >snd_pair + >snd_pair >snd_pair >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair + >fst_pair @le_plus + [@le_plus // cases (true_or_false (leb c1 c)) #Hcase + >Hcase normalize // @lt_to_le @not_le_to_lt @(leb_false_to_not_le ?? Hcase) + |@Hind + ] + ] + ] + |@O_plus [@O_plus_l //] @le_to_O #x + Hext // +qed. + +lemma constr_prim_rec: ∀s1,s2. constructible s1 → constructible s2 → + (∀n,r,m. 2 * r ≤ s2 〈n,〈r,m〉〉) → constructible (unary_pr s1 s2). +#s1 #s2 #Hs1 #Hs2 #Hincr @(CF_prim_rec … Hs1 Hs2) whd %{2} %{0} +#x #_ lapply (surj_pair x) * #a * #b #eqx >eqx whd in match (unary_pr ???); +>fst_pair >snd_pair +whd in match (unary_pr ???); >fst_pair >snd_pair elim a + [normalize // + |#n #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair + >prim_rec_S @transitive_le [| @(monotonic_le_plus_l … Hind)] + @transitive_le [| @(monotonic_le_plus_l … (Hincr n ? b))] + whd in match (unary_pr ???); >fst_pair >snd_pair // + ] +qed. + +(********************************* simulation *********************************) + +axiom sU : nat → nat. + +axiom monotonic_sU: ∀i1,i2,x1,x2,s1,s2. i1 ≤ i2 → x1 ≤ x2 → s1 ≤ s2 → + sU 〈i1,〈x1,s1〉〉 ≤ sU 〈i2,〈x2,s2〉〉. + +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) +#b1 * #c1 #eqy >eqy -eqy +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〉〉. + +definition pU_unary ≝ λp. pU (fst p) (fst (snd p)) (snd (snd p)). + +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. + + +(******************** complexity of g ********************) + +(*definition unary_g ≝ λh.λux. g h (fst ux) (snd ux). +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 (auxg h) → CF s (unary_g h). +#h #s #H1 @(CF_compS ? (auxg h) H1) +qed. + +definition aux1g ≝ + λh,ux. max_{i ∈[fst ux,snd ux[ | (λp. eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) 〈i,ux〉} + ((λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) 〈i,ux〉). + +lemma eq_aux : ∀h,x.aux1g h x = auxg h x. +#h #x @same_bigop + [#n #_ >fst_pair >snd_pair // |#n #_ #_ >fst_pair >snd_pair //] +qed. + +lemma compl_g2 : ∀h,s1,s2,s. + CF s1 + (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) → + CF s2 + (λ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 (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 (λ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 min_input_def whd in ⊢ (??%?); >minus_S_S @min_f_g #i #_ #_ +whd in ⊢ (??%%); >fst_pair >snd_pair // +qed. + +definition termb_aux ≝ λh. + termb_unary ∘ λp.〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉. + +lemma compl_g4 : ∀h,s1,s. + (CF s1 + (λ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 (λ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 + |@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //] +qed.*) + +(************************* a couple of technical lemmas ***********************) +lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0. +#a elim a // #n #Hind * + [#H @False_ind /2 by absurd/ | #b normalize #H @Hind @le_S_S_to_le /2/] +qed. + +lemma sigma_bound: ∀h,a,b. monotonic nat le h → + ∑_{i ∈ [a,S b[ }(h i) ≤ (S b-a)*h b. +#h #a #b #H cases (decidable_le a b) + [#leab cut (b = pred (S b - a + a)) + [Hb in match (h b); + generalize in match (S b -a); + #n elim n + [// + |#m #Hind >bigop_Strue [2://] @le_plus + [@H @le_n |@(transitive_le … Hind) @le_times [//] @H //] + ] + |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba + cut (S b -a = 0) [@minus_to_0 //] #Hcut >Hcut // + ] +qed. + +lemma sigma_bound_decr: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) → + ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a. +#h #a #b #H cases (decidable_le a b) + [#leab cut ((b -a) +a ≤ b) [/2 by le_minus_to_plus_r/] generalize in match (b -a); + #n elim n + [// + |#m #Hind >bigop_Strue [2://] #Hm + cut (m+a ≤ b) [@(transitive_le … Hm) //] #Hm1 + @le_plus [@H // |@(transitive_le … (Hind Hm1)) //] + ] + |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba + cut (b -a = 0) [@minus_to_0 @lt_to_le @ltba] #Hcut >Hcut // + ] +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〉)). +#s1 #Hs1 %{1} %{0} #n #_ >commutative_times minus_S_S //] +qed. + +lemma coroll2: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) → +O (λx.(snd x - fst x)*s1 〈fst x,x〉) (λx.∑_{i ∈[fst x,snd x[ }(s1 〈i,x〉)). +#s1 #Hs1 %{1} %{0} #n #_ >commutative_times fst_pair >snd_pair >fst_pair >snd_pair // ] +@(CF_comp … (λx.MSC 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 + [@O_plus + [@(O_trans … (λx.MSC (fst x) + MSC (max (fst (snd x)) (snd (snd x))))) + [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx + >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) + >distributive_times_plus @le_plus [//] + cases (surj_pair b) #c * #d #eqb >eqb + >fst_pair >snd_pair @(transitive_le … (MSC_pair …)) + whd in ⊢ (??%); @le_plus + [@monotonic_MSC @(le_maxl … (le_n …)) + |>commutative_times fst_pair >snd_pair // qed. + +lemma le_big : ∀x. x ≤ big x. +#x cases (surj_pair x) #a * #b #eqx >eqx @le_pair >fst_pair >snd_pair +[@(le_maxl … (le_n …)) | @(le_maxr … (le_n …))] +qed. + +definition faux2 ≝ λh. + (λx.MSC x + (snd (snd x)-fst x)* + (λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉). + +(*lemma compl_g7: ∀h. + constructible (λx. h (fst x) (snd x)) → + (∀n. monotonic ? le (h n)) → + CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉) + (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #hcostr #hmono @(monotonic_CF … (faux2 h)) + [#n normalize >fst_pair >snd_pair //] +@compl_g5 [2:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair +>fst_pair >snd_pair @monotonic_sU // @hmono @lexy +qed. + +lemma compl_g71: ∀h. + 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))〉〉) + (λp:ℕ.min_input h (fst p) (snd (snd p))). +#h #hcostr #hmono @(monotonic_CF … (compl_g7 h hcostr hmono)) #x +@le_plus [@monotonic_MSC //] +cases (decidable_le (fst x) (snd(snd x))) + [#Hle @le_times // @monotonic_sU + |#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt] + ] +qed.*) + +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 (MSC x)) >commutative_times @le_times + [// | @monotonic_MSC // ]] +@(O_trans … (coroll2 ??)) + [#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 [//|@(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 + [/2 by monotonic_le_minus_r/ + |@monotonic_sU // @hantimono [@le_S_S // |@ltb] + ] + |@monotonic_sU // @hantimono [@le_S_S // |@ltb] + ] + |@le_to_O #n >fst_pair >snd_pair + cut (max (fst n) n = n) [normalize >le_to_leb_true //] #Hmax >Hmax + >associative_plus >distributive_times_plus + @le_plus [@le_times [@le_S // |>big_def >Hmax //] |//] + ] +qed.*) + +definition sg ≝ λh,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)〉〉. + +lemma sg_def : ∀h,a,b. + sg h 〈a,b〉 = (S (b-a))*MSC 〈〈a,b〉,〈a,b〉〉 + + (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉. +#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair // +qed. + +(*lemma compl_g11 : ∀h. + 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) (unary_g h). +#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham) +qed.*) + +(**************************** closing the argument ****************************) + +let rec h_of_aux (r:nat →nat) (c,d,b:nat) on d : nat ≝ + match d with + [ O ⇒ c + | S d1 ⇒ (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + + d*(S d)*sU 〈〈b-d,b〉,〈b,r (h_of_aux r c d1 b)〉〉]. + +lemma h_of_aux_O: ∀r,c,b. + h_of_aux r c O b = c. +// qed. + +lemma h_of_aux_S : ∀r,c,d,b. + h_of_aux r c (S d) b = + (S (S d))*(MSC 〈〈b-(S d),b〉,〈b-(S d),b〉〉) + + (S d)*(S (S d))*sU 〈〈b-(S d),b〉,〈b,r(h_of_aux r c d b)〉〉. +// qed. + +lemma h_of_aux_prim_rec : ∀r,c,n,b. h_of_aux r c n b = + prim_rec (λx.c) + (λx.let d ≝ S(fst x) in + let b ≝ snd (snd x) in + (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) + + d*(S d)*sU 〈〈b-d,b〉,〈b,r (fst (snd x))〉〉) n b. +#r #c #n #b elim n + [>h_of_aux_O normalize // + |#n1 #Hind >h_of_aux_S >prim_rec_S >snd_pair >snd_pair >fst_pair + >fst_pair fst_pair >snd_pair >(minus_to_0 … Hle) // +qed. + +lemma h_of_def: ∀r,a,b.h_of r 〈a,b〉 = + let m ≝ max a b in + h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (b - a) b. +#r #a #b normalize >fst_pair >snd_pair // +qed. + +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. +#r #Hr #monor #d #d1 lapply d -d elim d1 + [#d #c #c1 #b #b1 #Hc #Hd @(le_n_O_elim ? Hd) #leb + >h_of_aux_O >h_of_aux_O // + |#m #Hind #d #c #c1 #b #b1 #lec #led #leb cases (le_to_or_lt_eq … led) + [#ltd @(transitive_le … (Hind … lec ? leb)) [@le_S_S_to_le @ltd] + >h_of_aux_S @(transitive_le ???? (le_plus_n …)) + >(times_n_1 (h_of_aux r c1 m b1)) in ⊢ (?%?); + >commutative_times @le_times [//|@(transitive_le … (Hr ?)) @sU_le] + |#Hd >Hd >h_of_aux_S >h_of_aux_S + cut (b-S m ≤ b1 - S m) [/2 by monotonic_le_minus_l/] #Hb1 + @le_plus [@le_times //] + [@monotonic_MSC @le_pair @le_pair // + |@le_times [//] @monotonic_sU + [@le_pair // |// |@monor @Hind //] + ] + ] + ] +qed. + +lemma mono_h_of2: ∀r.(∀x. x ≤ r x) → monotonic ? le r → + ∀i,b,b1. b ≤ b1 → h_of r 〈i,b〉 ≤ h_of r 〈i,b1〉. +#r #Hr #Hmono #i #a #b #leab >h_of_def >h_of_def +cut (max i a ≤ max i b) + [@to_max + [@(le_maxl … (le_n …))|@(transitive_le … leab) @(le_maxr … (le_n …))]] +#Hmax @(mono_h_of_aux r Hr Hmono) + [@monotonic_MSC @le_pair @le_pair @Hmax |/2 by monotonic_le_minus_l/ |@leab] +qed. + +axiom h_of_constr : ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + constructible (h_of r). + +(*lemma speed_compl: ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))). +#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …)) + [#x cases (surj_pair x) #a * #b #eqx >eqx + >sg_def cases (decidable_le b a) + [#leba >(minus_to_0 … leba) normalize in ⊢ (?%?); + h_of_def + cut (max a b = a) + [normalize cases (le_to_or_lt_eq … leba) + [#ltba >(lt_to_leb_false … ltba) % + |#eqba (le_to_leb_true … (le_n ?)) % ]] + #Hmax >Hmax normalize >(minus_to_0 … leba) normalize + @monotonic_MSC @le_pair @le_pair // + |#ltab >h_of_def >h_of_def + cut (max a b = b) + [normalize >(le_to_leb_true … ) [%] @lt_to_le @not_le_to_lt @ltab] + #Hmax >Hmax + 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) + [%{(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 + [@lt_to_le @not_le_to_lt // | //]] #eqd2 >eqd2 + normalize // + ] + |#n #a #b #leab #lebn >h_of_def >h_of_def + cut (max a n = n) + [normalize >le_to_leb_true [%|@(transitive_le … leab lebn)]] #Hmaxa + cut (max b n = n) + [normalize >(le_to_leb_true … lebn) %] #Hmaxb + >Hmaxa >Hmaxb @Hmono @(mono_h_of_aux r … Hr Hmono) // /2 by monotonic_le_minus_r/ + |#n #a #b #leab @Hmono @(mono_h_of2 … Hr Hmono … leab) + |@(constr_comp … Hconstr Hr) @(ext_constr (h_of r)) + [#x cases (surj_pair x) #a * #b #eqx >eqx >fst_pair >snd_pair //] + @(h_of_constr r Hr Hmono Hconstr) + ] +qed. + +lemma speed_compl_i: ∀r:nat →nat. + (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∀i. CF (λx.h_of r 〈i,x〉) (λx.g (λi,x. r(h_of r 〈i,x〉)) i x). +#r #Hr #Hmono #Hconstr #i +@(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 Hconstr)) #n // +qed.*) + +(**************************** the speedup theorem *****************************) +(*theorem pseudo_speedup: + ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∃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 #Hconstr +(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * +#Hcodei #HCi +(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} +(* sg is (λx.h_of r 〈i,x〉) *) +%{(λx. h_of r 〈S i,x〉)} +lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg +%[%[@condition_1 |@Hg] + |cases Hg #H1 * #j * #Hcodej #HCj + lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) + cases HCi #m * #a #Ha %{m} %{(max (S i) a)} #n #ltin @lt_to_le @not_le_to_lt + @(not_to_not … Hcond2) -Hcond2 #Hlesf %{n} % + [@(transitive_le … ltin) @(le_maxl … (le_n …))] + cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] + #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) // + ] +qed. + +theorem pseudo_speedup': + ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r → + ∃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 #Hconstr +(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #H * #i * +#Hcodei #HCi +(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *) +%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))} +(* sg is (λx.h_of r 〈i,x〉) *) +%{(λx. h_of r 〈S i,x〉)} +lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg +%[%[@condition_1 |@Hg] + |cases Hg #H1 * #j * #Hcodej #HCj + lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *) + cases HCi #m * #a #Ha + %{m} %{(max (S i) a)} #n #ltin @not_le_to_lt @(not_to_not … Hcond2) -Hcond2 #Hlesf + %{n} % [@(transitive_le … ltin) @(le_maxl … (le_n …))] + cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))] + #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) + @Hmono @(mono_h_of2 … Hr Hmono … ltin) + ] +qed.*) + \ No newline at end of file