From: Claudio Sacerdoti Coen Date: Wed, 26 Dec 2018 15:09:34 +0000 (+0100) Subject: Broken libs moved to broken_lib X-Git-Tag: make_still_working~229^2~19 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=2e481fd54d89b2a48f6f38f9562125136aa41d83 Broken libs moved to broken_lib Some can probably be restored with some love. --- diff --git a/matita/matita/broken_lib/binding/db.ma b/matita/matita/broken_lib/binding/db.ma new file mode 100644 index 000000000..534874f21 --- /dev/null +++ b/matita/matita/broken_lib/binding/db.ma @@ -0,0 +1,110 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basics/lists/list.ma". + +axiom alpha : Type[0]. +notation "𝔸" non associative with precedence 90 for @{'alphabet}. +interpretation "set of names" 'alphabet = alpha. + +inductive tp : Type[0] ≝ +| top : tp +| arr : tp → tp → tp. +inductive tm : Type[0] ≝ +| var : nat → tm +| par : 𝔸 → tm +| abs : tp → tm → tm +| app : tm → tm → tm. + +let rec Nth T n (l:list T) on n ≝ + match l with + [ nil ⇒ None ? + | cons hd tl ⇒ match n with + [ O ⇒ Some ? hd + | S n0 ⇒ Nth T n0 tl ] ]. + +inductive judg : list tp → tm → tp → Prop ≝ +| t_var : ∀g,n,t.Nth ? n g = Some ? t → judg g (var n) t +| t_app : ∀g,m,n,t,u.judg g m (arr t u) → judg g n t → judg g (app m n) u +| t_abs : ∀g,t,m,u.judg (t::g) m u → judg g (abs t m) (arr t u). + +definition Env := list (𝔸 × tp). + +axiom vclose_env : Env → list tp. +axiom vclose_tm : Env → tm → tm. +axiom Lam : 𝔸 → tp → tm → tm. +definition Judg ≝ λG,M,T.judg (vclose_env G) (vclose_tm G M) T. +definition dom ≝ λG:Env.map ?? (fst ??) G. + +definition sctx ≝ 𝔸 × tm. +axiom swap_tm : 𝔸 → 𝔸 → tm → tm. +definition sctx_app : sctx → 𝔸 → tm ≝ λM0,Y.let 〈X,M〉 ≝ M0 in swap_tm X Y M. + +axiom in_list : ∀A:Type[0].A → list A → Prop. +interpretation "list membership" 'mem x l = (in_list ? x l). +interpretation "list non-membership" 'notmem x l = (Not (in_list ? x l)). + +axiom in_Env : 𝔸 × tp → Env → Prop. +notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}. +interpretation "Env membership" 'lefttriangle x l = (in_Env x l). + +let rec FV M ≝ match M with + [ par X ⇒ [X] + | app M1 M2 ⇒ FV M1@FV M2 + | abs T M0 ⇒ FV M0 + | _ ⇒ [ ] ]. + +(* axiom Lookup : 𝔸 → Env → option tp. *) + +(* forma alto livello del judgment + t_abs* : ∀G,T,X,M,U. + (∀Y ∉ supp(M).Judg (〈Y,T〉::G) (M[Y]) U) → + Judg G (Lam X T (M[X])) (arr T U) *) + +(* prima dimostrare, poi perfezionare gli assiomi, poi dimostrarli *) + +axiom Judg_ind : ∀P:Env → tm → tp → Prop. + (∀X,G,T.〈X,T〉 ◃ G → P G (par X) T) → + (∀G,M,N,T,U. + Judg G M (arr T U) → Judg G N T → + P G M (arr T U) → P G N T → P G (app M N) U) → + (∀G,T1,T2,X,M1. + (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → Judg (〈Y,T1〉::G) (sctx_app M1 Y) T2) → + (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → P (〈Y,T1〉::G) (sctx_app M1 Y) T2) → + P G (Lam X T1 (sctx_app M1 X)) (arr T1 T2)) → + ∀G,M,T.Judg G M T → P G M T. + +axiom t_par : ∀X,G,T.〈X,T〉 ◃ G → Judg G (par X) T. +axiom t_app2 : ∀G,M,N,T,U.Judg G M (arr T U) → Judg G N T → Judg G (app M N) U. +axiom t_Lam : ∀G,X,M,T,U.Judg (〈X,T〉::G) M U → Judg G (Lam X T M) (arr T U). + +definition subenv ≝ λG1,G2.∀x.x ◃ G1 → x ◃ G2. +interpretation "subenv" 'subseteq G1 G2 = (subenv G1 G2). + +axiom daemon : ∀P:Prop.P. + +theorem weakening : ∀G1,G2,M,T.G1 ⊆ G2 → Judg G1 M T → Judg G2 M T. +#G1 #G2 #M #T #Hsub #HJ lapply Hsub lapply G2 -G2 change with (∀G2.?) +@(Judg_ind … HJ) +[ #X #G #T0 #Hin #G2 #Hsub @t_par @Hsub // +| #G #M0 #N #T0 #U #HM0 #HN #IH1 #IH2 #G2 #Hsub @t_app2 + [| @IH1 // | @IH2 // ] +| #G #T1 #T2 #X #M1 #HM1 #IH #G2 #Hsub @t_Lam @IH + [ (* trivial property of Lam *) @daemon + | (* trivial property of subenv *) @daemon ] +] +qed. + +(* Serve un tipo Tm per i termini localmente chiusi e i suoi principi di induzione e + ricorsione *) \ No newline at end of file diff --git a/matita/matita/broken_lib/binding/fp.ma b/matita/matita/broken_lib/binding/fp.ma new file mode 100644 index 000000000..14ba2f774 --- /dev/null +++ b/matita/matita/broken_lib/binding/fp.ma @@ -0,0 +1,296 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "binding/names.ma". + +(* permutations *) +definition finite_perm : ∀X:Nset.(X → X) → Prop ≝ + λX,f.injective X X f ∧ surjective X X f ∧ ∃l.∀x.x ∉ l → f x = x. + +(* maps a permutation to a list of parameters *) +definition Pi_list : ∀X:Nset.(X → X) → list X → list X ≝ + λX,p,xl.map ?? (λx.p x) xl. + +interpretation "permutation of X list" 'middot p x = (Pi_list p x). + +definition swap : ∀N:Nset.N → N → N → N ≝ + λN,u,v,x.match (x == u) with + [true ⇒ v + |false ⇒ match (x == v) with + [true ⇒ u + |false ⇒ x]]. + +axiom swap_right : ∀N,x,y.swap N x y y = x. +(* +#N x y;nnormalize;nrewrite > (p_eqb3 ? y y …);//; +nlapply (refl ? (y ≟ x));ncases (y ≟ x) in ⊢ (???% → %);nnormalize;//; +#H1;napply p_eqb1;//; +nqed. +*) + +axiom swap_left : ∀N,x,y.swap N x y x = y. +(* +#N x y;nnormalize;nrewrite > (p_eqb3 ? x x …);//; +nqed. +*) + +axiom swap_other : ∀N,x,y,z.x ≠ z → y ≠ z → swap N x y z = z. +(* +#N x y z H1 H2;nnormalize;nrewrite > (p_eqb4 …); +##[nrewrite > (p_eqb4 …);//;@;ncases H2;/2/; +##|@;ncases H1;/2/ +##] +nqed. +*) + +axiom swap_inv : ∀N,x,y,z.swap N x y (swap N x y z) = z. +(* +#N x y z;nlapply (refl ? (x ≟ z));ncases (x ≟ z) in ⊢ (???% → ?);#H1 +##[nrewrite > (p_eqb1 … H1);nrewrite > (swap_left …);//; +##|nlapply (refl ? (y ≟ z));ncases (y ≟ z) in ⊢ (???% → ?);#H2 + ##[nrewrite > (p_eqb1 … H2);nrewrite > (swap_right …);//; + ##|nrewrite > (swap_other …) in ⊢ (??(????%)?);/2/; + nrewrite > (swap_other …);/2/; + ##] +##] +nqed. +*) + +axiom swap_fp : ∀N,x1,x2.finite_perm ? (swap N x1 x2). +(* +#N x1 x2;@ +##[@ + ##[nwhd;#xa xb;nnormalize;nlapply (refl ? (xa ≟ x1)); + ncases (xa ≟ x1) in ⊢ (???% → %);#H1 + ##[nrewrite > (p_eqb1 … H1);nlapply (refl ? (xb ≟ x1)); + ncases (xb ≟ x1) in ⊢ (???% → %);#H2 + ##[nrewrite > (p_eqb1 … H2);// + ##|nlapply (refl ? (xb ≟ x2)); + ncases (xb ≟ x2) in ⊢ (???% → %);#H3 + ##[nnormalize;#H4;nrewrite > H4 in H3; + #H3;nrewrite > H3 in H2;#H2;ndestruct (H2) + ##|nnormalize;#H4;nrewrite > H4 in H3; + nrewrite > (p_eqb3 …);//;#H5;ndestruct (H5) + ##] + ##] + ##|nlapply (refl ? (xa ≟ x2)); + ncases (xa ≟ x2) in ⊢ (???% → %);#H2 + ##[nrewrite > (p_eqb1 … H2);nlapply (refl ? (xb ≟ x1)); + ncases (xb ≟ x1) in ⊢ (???% → %);#H3 + ##[nnormalize;#H4;nrewrite > H4 in H3; + #H3;nrewrite > (p_eqb1 … H3);@ + ##|nnormalize;nlapply (refl ? (xb ≟ x2)); + ncases (xb ≟ x2) in ⊢ (???% → %);#H4 + ##[nrewrite > (p_eqb1 … H4);// + ##|nnormalize;#H5;nrewrite > H5 in H3; + nrewrite > (p_eqb3 …);//;#H6;ndestruct (H6); + ##] + ##] + ##|nnormalize;nlapply (refl ? (xb ≟ x1)); + ncases (xb ≟ x1) in ⊢ (???% → %);#H3 + ##[nnormalize;#H4;nrewrite > H4 in H2;nrewrite > (p_eqb3 …);//; + #H5;ndestruct (H5) + ##|nlapply (refl ? (xb ≟ x2)); + ncases (xb ≟ x2) in ⊢ (???% → %);#H4 + ##[nnormalize;#H5;nrewrite > H5 in H1;nrewrite > (p_eqb3 …);//; + #H6;ndestruct (H6) + ##|nnormalize;// + ##] + ##] + ##] + ##] + ##|nwhd;#z;nnormalize;nlapply (refl ? (z ≟ x1)); + ncases (z ≟ x1) in ⊢ (???% → %);#H1 + ##[nlapply (refl ? (z ≟ x2)); + ncases (z ≟ x2) in ⊢ (???% → %);#H2 + ##[@ z;nrewrite > H1;nrewrite > H2;napply p_eqb1;// + ##|@ x2;nrewrite > (p_eqb4 …); + ##[nrewrite > (p_eqb3 …);//; + nnormalize;napply p_eqb1;// + ##|nrewrite < (p_eqb1 … H1);@;#H3;nrewrite > H3 in H2; + nrewrite > (p_eqb3 …);//;#H2;ndestruct (H2) + ##] + ##] + ##|nlapply (refl ? (z ≟ x2)); + ncases (z ≟ x2) in ⊢ (???% → %);#H2 + ##[@ x1;nrewrite > (p_eqb3 …);//; + napply p_eqb1;nnormalize;// + ##|@ z;nrewrite > H1;nrewrite > H2;@; + ##] + ##] + ##] +##|@ [x1;x2];#x0 H1;nrewrite > (swap_other …) + ##[@ + ##|@;#H2;nrewrite > H2 in H1;*;#H3;napply H3;/2/; + ##|@;#H2;nrewrite > H2 in H1;*;#H3;napply H3;//; + ##] +##] +nqed. +*) + +axiom inj_swap : ∀N,u,v.injective ?? (swap N u v). +(* +#N u v;ncases (swap_fp N u v);*;#H1 H2 H3;//; +nqed. +*) + +axiom surj_swap : ∀N,u,v.surjective ?? (swap N u v). +(* +#N u v;ncases (swap_fp N u v);*;#H1 H2 H3;//; +nqed. +*) + +axiom finite_swap : ∀N,u,v.∃l.∀x.x ∉ l → swap N u v x = x. +(* +#N u v;ncases (swap_fp N u v);*;#H1 H2 H3;//; +nqed. +*) + +(* swaps two lists of parameters +definition Pi_swap_list : ∀xl,xl':list X.X → X ≝ + λxl,xl',x.foldr2 ??? (λu,v,r.swap ? u v r) x xl xl'. + +nlemma fp_swap_list : + ∀xl,xl'.finite_perm ? (Pi_swap_list xl xl'). +#xl xl';@ +##[@; + ##[ngeneralize in match xl';nelim xl + ##[nnormalize;//; + ##|#x0 xl0;#IH xl'';nelim xl'' + ##[nnormalize;// + ##|#x1 xl1 IH1 y0 y1;nchange in ⊢ (??%% → ?) with (swap ????); + #H1;nlapply (inj_swap … H1);#H2; + nlapply (IH … H2);// + ##] + ##] + ##|ngeneralize in match xl';nelim xl + ##[nnormalize;#_;#z;@z;@ + ##|#x' xl0 IH xl'';nelim xl'' + ##[nnormalize;#z;@z;@ + ##|#x1 xl1 IH1 z; + nchange in ⊢ (??(λ_.???%)) with (swap ????); + ncases (surj_swap X x' x1 z);#x2 H1; + ncases (IH xl1 x2);#x3 H2;@ x3; + nrewrite < H2;napply H1 + ##] + ##] + ##] +##|ngeneralize in match xl';nelim xl + ##[#;@ [];#;@ + ##|#x0 xl0 IH xl'';ncases xl'' + ##[@ [];#;@ + ##|#x1 xl1;ncases (IH xl1);#xl2 H1; + ncases (finite_swap X x0 x1);#xl3 H2; + @ (xl2@xl3);#x2 H3; + nchange in ⊢ (??%?) with (swap ????); + nrewrite > (H1 …); + ##[nrewrite > (H2 …);//;@;#H4;ncases H3;#H5;napply H5; + napply in_list_to_in_list_append_r;// + ##|@;#H4;ncases H3;#H5;napply H5; + napply in_list_to_in_list_append_l;// + ##] + ##] + ##] +##] +nqed. + +(* the 'reverse' swap of lists of parameters + composing Pi_swap_list and Pi_swap_list_r yields the identity function + (see the Pi_swap_list_inv lemma) *) +ndefinition Pi_swap_list_r : ∀xl,xl':list X. Pi ≝ + λxl,xl',x.foldl2 ??? (λr,u,v.swap ? u v r ) x xl xl'. + +nlemma fp_swap_list_r : + ∀xl,xl'.finite_perm ? (Pi_swap_list_r xl xl'). +#xl xl';@ +##[@; + ##[ngeneralize in match xl';nelim xl + ##[nnormalize;//; + ##|#x0 xl0;#IH xl'';nelim xl'' + ##[nnormalize;// + ##|#x1 xl1 IH1 y0 y1;nwhd in ⊢ (??%% → ?); + #H1;nlapply (IH … H1);#H2; + napply (inj_swap … H2); + ##] + ##] + ##|ngeneralize in match xl';nelim xl + ##[nnormalize;#_;#z;@z;@ + ##|#x' xl0 IH xl'';nelim xl'' + ##[nnormalize;#z;@z;@ + ##|#x1 xl1 IH1 z;nwhd in ⊢ (??(λ_.???%)); + ncases (IH xl1 z);#x2 H1; + ncases (surj_swap X x' x1 x2);#x3 H2; + @ x3;nrewrite < H2;napply H1; + ##] + ##] + ##] +##|ngeneralize in match xl';nelim xl + ##[#;@ [];#;@ + ##|#x0 xl0 IH xl'';ncases xl'' + ##[@ [];#;@ + ##|#x1 xl1; + ncases (IH xl1);#xl2 H1; + ncases (finite_swap X x0 x1);#xl3 H2; + @ (xl2@xl3);#x2 H3;nwhd in ⊢ (??%?); + nrewrite > (H2 …); + ##[nrewrite > (H1 …);//;@;#H4;ncases H3;#H5;napply H5; + napply in_list_to_in_list_append_l;// + ##|@;#H4;ncases H3;#H5;napply H5; + napply in_list_to_in_list_append_r;// + ##] + ##] + ##] +##] +nqed. + +nlemma Pi_swap_list_inv : + ∀xl1,xl2,x. + Pi_swap_list xl1 xl2 (Pi_swap_list_r xl1 xl2 x) = x. +#xl;nelim xl +##[#;@ +##|#x1 xl1 IH xl';ncases xl' + ##[#;@ + ##|#x2 xl2;#x; + nchange in ⊢ (??%?) with + (swap ??? (Pi_swap_list ?? + (Pi_swap_list_r ?? (swap ????)))); + nrewrite > (IH xl2 ?);napply swap_inv; + ##] +##] +nqed. + +nlemma Pi_swap_list_fresh : + ∀x,xl1,xl2.x ∉ xl1 → x ∉ xl2 → Pi_swap_list xl1 xl2 x = x. +#x xl1;nelim xl1 +##[#;@ +##|#x3 xl3 IH xl2' H1;ncases xl2' + ##[#;@ + ##|#x4 xl4 H2;ncut (x ∉ xl3 ∧ x ∉ xl4); + ##[@ + ##[@;#H3;ncases H1;#H4;napply H4;/2/; + ##|@;#H3;ncases H2;#H4;napply H4;/2/ + ##] + ##] *; #H1' H2'; + nchange in ⊢ (??%?) with (swap ????); + nrewrite > (swap_other …) + ##[napply IH;//; + ##|nchange in ⊢ (?(???%)) with (Pi_swap_list ???); + nrewrite > (IH …);//;@;#H3;ncases H2;#H4;napply H4;//; + ##|nchange in ⊢ (?(???%)) with (Pi_swap_list ???); + nrewrite > (IH …);//;@;#H3;ncases H1;#H4;napply H4;// + ##] + ##] +##] +nqed. +*) \ No newline at end of file diff --git a/matita/matita/broken_lib/binding/ln.ma b/matita/matita/broken_lib/binding/ln.ma new file mode 100644 index 000000000..ca7f574ed --- /dev/null +++ b/matita/matita/broken_lib/binding/ln.ma @@ -0,0 +1,716 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basics/lists/list.ma". +include "basics/deqsets.ma". +include "binding/names.ma". +include "binding/fp.ma". + +axiom alpha : Nset. +notation "𝔸" non associative with precedence 90 for @{'alphabet}. +interpretation "set of names" 'alphabet = alpha. + +inductive tp : Type[0] ≝ +| top : tp +| arr : tp → tp → tp. +inductive pretm : Type[0] ≝ +| var : nat → pretm +| par : 𝔸 → pretm +| abs : tp → pretm → pretm +| app : pretm → pretm → pretm. + +let rec Nth T n (l:list T) on n ≝ + match l with + [ nil ⇒ None ? + | cons hd tl ⇒ match n with + [ O ⇒ Some ? hd + | S n0 ⇒ Nth T n0 tl ] ]. + +let rec vclose_tm_aux u x k ≝ match u with + [ var n ⇒ if (leb k n) then var (S n) else u + | par x0 ⇒ if (x0 == x) then (var k) else u + | app v1 v2 ⇒ app (vclose_tm_aux v1 x k) (vclose_tm_aux v2 x k) + | abs s v ⇒ abs s (vclose_tm_aux v x (S k)) ]. +definition vclose_tm ≝ λu,x.vclose_tm_aux u x O. + +definition vopen_var ≝ λn,x,k.match eqb n k with + [ true ⇒ par x + | false ⇒ match leb n k with + [ true ⇒ var n + | false ⇒ var (pred n) ] ]. + +let rec vopen_tm_aux u x k ≝ match u with + [ var n ⇒ vopen_var n x k + | par x0 ⇒ u + | app v1 v2 ⇒ app (vopen_tm_aux v1 x k) (vopen_tm_aux v2 x k) + | abs s v ⇒ abs s (vopen_tm_aux v x (S k)) ]. +definition vopen_tm ≝ λu,x.vopen_tm_aux u x O. + +let rec FV u ≝ match u with + [ par x ⇒ [x] + | app v1 v2 ⇒ FV v1@FV v2 + | abs s v ⇒ FV v + | _ ⇒ [ ] ]. + +definition lam ≝ λx,s,u.abs s (vclose_tm u x). + +let rec Pi_map_tm p u on u ≝ match u with +[ par x ⇒ par (p x) +| var _ ⇒ u +| app v1 v2 ⇒ app (Pi_map_tm p v1) (Pi_map_tm p v2) +| abs s v ⇒ abs s (Pi_map_tm p v) ]. + +interpretation "permutation of tm" 'middot p x = (Pi_map_tm p x). + +notation "hvbox(u⌈x⌉)" + with precedence 45 + for @{ 'open $u $x }. + +(* +notation "hvbox(u⌈x⌉)" + with precedence 45 + for @{ 'open $u $x }. +notation "❴ u ❵ x" non associative with precedence 90 for @{ 'open $u $x }. +*) +interpretation "ln term variable open" 'open u x = (vopen_tm u x). +notation < "hvbox(ν x break . u)" + with precedence 20 +for @{'nu $x $u }. +notation > "ν list1 x sep , . term 19 u" with precedence 20 + for ${ fold right @{$u} rec acc @{'nu $x $acc)} }. +interpretation "ln term variable close" 'nu x u = (vclose_tm u x). + +let rec tm_height u ≝ match u with +[ app v1 v2 ⇒ S (max (tm_height v1) (tm_height v2)) +| abs s v ⇒ S (tm_height v) +| _ ⇒ O ]. + +theorem le_n_O_rect_Type0 : ∀n:nat. n ≤ O → ∀P: nat →Type[0]. P O → P n. +#n (cases n) // #a #abs cases (?:False) /2/ qed. + +theorem nat_rect_Type0_1 : ∀n:nat.∀P:nat → Type[0]. +(∀m.(∀p. p < m → P p) → P m) → P n. +#n #P #H +cut (∀q:nat. q ≤ n → P q) /2/ +(elim n) + [#q #HleO (* applica male *) + @(le_n_O_rect_Type0 ? HleO) + @H #p #ltpO cases (?:False) /2/ (* 3 *) + |#p #Hind #q #HleS + @H #a #lta @Hind @le_S_S_to_le /2/ + ] +qed. + +lemma leb_false_to_lt : ∀n,m. leb n m = false → m < n. +#n elim n +[ #m normalize #H destruct(H) +| #n0 #IH * // #m normalize #H @le_S_S @IH // ] +qed. + +lemma nominal_eta_aux : ∀x,u.x ∉ FV u → ∀k.vclose_tm_aux (vopen_tm_aux u x k) x k = u. +#x #u elim u +[ #n #_ #k normalize cases (decidable_eq_nat n k) #Hnk + [ >Hnk >eqb_n_n normalize >(\b ?) // + | >(not_eq_to_eqb_false … Hnk) normalize cases (true_or_false (leb n k)) #Hleb + [ >Hleb normalize >(?:leb k n = false) // + @lt_to_leb_false @not_eq_to_le_to_lt /2/ + | >Hleb normalize >(?:leb k (pred n) = true) normalize + [ cases (leb_false_to_lt … Hleb) // + | @le_to_leb_true cases (leb_false_to_lt … Hleb) normalize /2/ ] ] ] +| #y normalize #Hy >(\bf ?) // @(not_to_not … Hy) // +| #s #v #IH normalize #Hv #k >IH // @Hv +| #v1 #v2 #IH1 #IH2 normalize #Hv1v2 #k + >IH1 [ >IH2 // | @(not_to_not … Hv1v2) @in_list_to_in_list_append_l ] + @(not_to_not … Hv1v2) @in_list_to_in_list_append_r ] +qed. + +corollary nominal_eta : ∀x,u.x ∉ FV u → (νx.u⌈x⌉) = u. +#x #u #Hu @nominal_eta_aux // +qed. + +lemma eq_height_vopen_aux : ∀v,x,k.tm_height (vopen_tm_aux v x k) = tm_height v. +#v #x elim v +[ #n #k normalize cases (eqb n k) // cases (leb n k) // +| #u #k % +| #s #u #IH #k normalize >IH % +| #u1 #u2 #IH1 #IH2 #k normalize >IH1 >IH2 % ] +qed. + +corollary eq_height_vopen : ∀v,x.tm_height (v⌈x⌉) = tm_height v. +#v #x @eq_height_vopen_aux +qed. + +theorem pretm_ind_plus_aux : + ∀P:pretm → Type[0]. + (∀x:𝔸.P (par x)) → + (∀n:ℕ.P (var n)) → + (∀v1,v2. P v1 → P v2 → P (app v1 v2)) → + ∀C:list 𝔸. + (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) → + ∀n,u.tm_height u ≤ n → P u. +#P #Hpar #Hvar #Happ #C #Hlam #n change with ((λn.?) n); @(nat_rect_Type0_1 n ??) +#m cases m +[ #_ * /2/ + [ normalize #s #v #Hfalse cases (?:False) cases (not_le_Sn_O (tm_height v)) /2/ + | #v1 #v2 whd in ⊢ (?%?→?); #Hfalse cases (?:False) cases (not_le_Sn_O (max ??)) + [ #H @H @Hfalse|*:skip] ] ] +-m #m #IH * /2/ +[ #s #v whd in ⊢ (?%?→?); #Hv + lapply (p_fresh … (C@FV v)) letin y ≝ (N_fresh … (C@FV v)) #Hy + >(?:abs s v = lam y s (v⌈y⌉)) + [| whd in ⊢ (???%); >nominal_eta // @(not_to_not … Hy) @in_list_to_in_list_append_r ] + @Hlam + [ @(not_to_not … Hy) @in_list_to_in_list_append_r + | @(not_to_not … Hy) @in_list_to_in_list_append_l ] + @IH [| @Hv | >eq_height_vopen % ] +| #v1 #v2 whd in ⊢ (?%?→?); #Hv @Happ + [ @IH [| @Hv | @le_max_1 ] | @IH [| @Hv | @le_max_2 ] ] ] +qed. + +corollary pretm_ind_plus : + ∀P:pretm → Type[0]. + (∀x:𝔸.P (par x)) → + (∀n:ℕ.P (var n)) → + (∀v1,v2. P v1 → P v2 → P (app v1 v2)) → + ∀C:list 𝔸. + (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) → + ∀u.P u. +#P #Hpar #Hvar #Happ #C #Hlam #u @pretm_ind_plus_aux /2/ +qed. + +(* maps a permutation to a list of terms *) +definition Pi_map_list : (𝔸 → 𝔸) → list 𝔸 → list 𝔸 ≝ map 𝔸 𝔸 . + +(* interpretation "permutation of name list" 'middot p x = (Pi_map_list p x).*) + +(* +inductive tm : pretm → Prop ≝ +| tm_par : ∀x:𝔸.tm (par x) +| tm_app : ∀u,v.tm u → tm v → tm (app u v) +| tm_lam : ∀x,s,u.tm u → tm (lam x s u). + +inductive ctx_aux : nat → pretm → Prop ≝ +| ctx_var : ∀n,k.n < k → ctx_aux k (var n) +| ctx_par : ∀x,k.ctx_aux k (par x) +| ctx_app : ∀u,v,k.ctx_aux k u → ctx_aux k v → ctx_aux k (app u v) +(* è sostituibile da ctx_lam ? *) +| ctx_abs : ∀s,u.ctx_aux (S k) u → ctx_aux k (abs s u). +*) + +inductive tm_or_ctx (k:nat) : pretm → Type[0] ≝ +| toc_var : ∀n.n < k → tm_or_ctx k (var n) +| toc_par : ∀x.tm_or_ctx k (par x) +| toc_app : ∀u,v.tm_or_ctx k u → tm_or_ctx k v → tm_or_ctx k (app u v) +| toc_lam : ∀x,s,u.tm_or_ctx k u → tm_or_ctx k (lam x s u). + +definition tm ≝ λt.tm_or_ctx O t. +definition ctx ≝ λt.tm_or_ctx 1 t. + +definition check_tm ≝ λu,k. + pretm_ind_plus ? + (λ_.true) + (λn.leb (S n) k) + (λv1,v2,rv1,rv2.rv1 ∧ rv2) + [] (λx,s,v,px,pC,rv.rv) + u. + +axiom pretm_ind_plus_app : ∀P,u,v,C,H1,H2,H3,H4. + pretm_ind_plus P H1 H2 H3 C H4 (app u v) = + H3 u v (pretm_ind_plus P H1 H2 H3 C H4 u) (pretm_ind_plus P H1 H2 H3 C H4 v). + +axiom pretm_ind_plus_lam : ∀P,x,s,u,C,px,pC,H1,H2,H3,H4. + pretm_ind_plus P H1 H2 H3 C H4 (lam x s (u⌈x⌉)) = + H4 x s u px pC (pretm_ind_plus P H1 H2 H3 C H4 (u⌈x⌉)). + +record TM : Type[0] ≝ { + pretm_of_TM :> pretm; + tm_of_TM : check_tm pretm_of_TM O = true +}. + +record CTX : Type[0] ≝ { + pretm_of_CTX :> pretm; + ctx_of_CTX : check_tm pretm_of_CTX 1 = true +}. + +inductive tm2 : pretm → Type[0] ≝ +| tm_par : ∀x.tm2 (par x) +| tm_app : ∀u,v.tm2 u → tm2 v → tm2 (app u v) +| tm_lam : ∀x,s,u.x ∉ FV u → (∀y.y ∉ FV u → tm2 (u⌈y⌉)) → tm2 (lam x s (u⌈x⌉)). + +(* +inductive tm' : pretm → Prop ≝ +| tm_par : ∀x.tm' (par x) +| tm_app : ∀u,v.tm' u → tm' v → tm' (app u v) +| tm_lam : ∀x,s,u,C.x ∉ FV u → x ∉ C → (∀y.y ∉ FV u → tm' (❴u❵y)) → tm' (lam x s (❴u❵x)). +*) + +lemma pi_vclose_tm : + ∀z1,z2,x,u.swap 𝔸 z1 z2·(νx.u) = (ν swap ? z1 z2 x.swap 𝔸 z1 z2 · u). +#z1 #z2 #x #u +change with (vclose_tm_aux ???) in match (vclose_tm ??); +change with (vclose_tm_aux ???) in ⊢ (???%); lapply O elim u normalize // +[ #n #k cases (leb k n) normalize % +| #x0 #k cases (true_or_false (x0==z1)) #H1 >H1 normalize + [ cases (true_or_false (x0==x)) #H2 >H2 normalize + [ <(\P H2) >H1 normalize >(\b (refl ? z2)) % + | >H1 normalize cases (true_or_false (x==z1)) #H3 >H3 normalize + [ >(\P H3) in H2; >H1 #Hfalse destruct (Hfalse) + | cases (true_or_false (x==z2)) #H4 >H4 normalize + [ cases (true_or_false (z2==z1)) #H5 >H5 normalize // + >(\P H5) in H4; >H3 #Hfalse destruct (Hfalse) + | >(\bf ?) // @sym_not_eq @(\Pf H4) ] + ] + ] + | cases (true_or_false (x0==x)) #H2 >H2 normalize + [ <(\P H2) >H1 normalize >(\b (refl ??)) % + | >H1 normalize cases (true_or_false (x==z1)) #H3 >H3 normalize + [ cases (true_or_false (x0==z2)) #H4 >H4 normalize + [ cases (true_or_false (z1==z2)) #H5 >H5 normalize // + <(\P H5) in H4; <(\P H3) >H2 #Hfalse destruct (Hfalse) + | >H4 % ] + | cases (true_or_false (x0==z2)) #H4 >H4 normalize + [ cases (true_or_false (x==z2)) #H5 >H5 normalize + [ <(\P H5) in H4; >H2 #Hfalse destruct (Hfalse) + | >(\bf ?) // @sym_not_eq @(\Pf H3) ] + | cases (true_or_false (x==z2)) #H5 >H5 normalize + [ >H1 % + | >H2 % ] + ] + ] + ] + ] +] +qed. + +lemma pi_vopen_tm : + ∀z1,z2,x,u.swap 𝔸 z1 z2·(u⌈x⌉) = (swap 𝔸 z1 z2 · u⌈swap 𝔸 z1 z2 x⌉). +#z1 #z2 #x #u +change with (vopen_tm_aux ???) in match (vopen_tm ??); +change with (vopen_tm_aux ???) in ⊢ (???%); lapply O elim u normalize // +#n #k cases (true_or_false (eqb n k)) #H1 >H1 normalize // +cases (true_or_false (leb n k)) #H2 >H2 normalize // +qed. + +lemma pi_lam : + ∀z1,z2,x,s,u.swap 𝔸 z1 z2 · lam x s u = lam (swap 𝔸 z1 z2 x) s (swap 𝔸 z1 z2 · u). +#z1 #z2 #x #s #u whd in ⊢ (???%); <(pi_vclose_tm …) % +qed. + +lemma eqv_FV : ∀z1,z2,u.FV (swap 𝔸 z1 z2 · u) = Pi_map_list (swap 𝔸 z1 z2) (FV u). +#z1 #z2 #u elim u // +[ #s #v normalize // +| #v1 #v2 normalize /2/ ] +qed. + +lemma swap_inv_tm : ∀z1,z2,u.swap 𝔸 z1 z2 · (swap 𝔸 z1 z2 · u) = u. +#z1 #z2 #u elim u [1,3,4:normalize //] +#x whd in ⊢ (??%?); >swap_inv % +qed. + +lemma eqv_in_list : ∀x,l,z1,z2.x ∈ l → swap 𝔸 z1 z2 x ∈ Pi_map_list (swap 𝔸 z1 z2) l. +#x #l #z1 #z2 #Hin elim Hin +[ #x0 #l0 % +| #x1 #x2 #l0 #Hin #IH %2 @IH ] +qed. + +lemma eqv_tm2 : ∀u.tm2 u → ∀z1,z2.tm2 ((swap ? z1 z2)·u). +#u #Hu #z1 #z2 letin p ≝ (swap ? z1 z2) elim Hu /2/ +#x #s #v #Hx #Hv #IH >pi_lam >pi_vopen_tm %3 +[ @(not_to_not … Hx) -Hx #Hx + <(swap_inv ? z1 z2 x) <(swap_inv_tm z1 z2 v) >eqv_FV @eqv_in_list // +| #y #Hy <(swap_inv ? z1 z2 y) + eqv_FV @eqv_in_list // +] +qed. + +lemma vclose_vopen_aux : ∀x,u,k.vopen_tm_aux (vclose_tm_aux u x k) x k = u. +#x #u elim u normalize // +[ #n #k cases (true_or_false (leb k n)) #H >H whd in ⊢ (??%?); + [ cases (true_or_false (eqb (S n) k)) #H1 >H1 + [ <(eqb_true_to_eq … H1) in H; #H lapply (leb_true_to_le … H) -H #H + cases (le_to_not_lt … H) -H #H cases (H ?) % + | whd in ⊢ (??%?); >lt_to_leb_false // @le_S_S /2/ ] + | cases (true_or_false (eqb n k)) #H1 >H1 normalize + [ >(eqb_true_to_eq … H1) in H; #H lapply (leb_false_to_not_le … H) -H + * #H cases (H ?) % + | >le_to_leb_true // @not_lt_to_le % #H2 >le_to_leb_true in H; + [ #H destruct (H) | /2/ ] + ] + ] +| #x0 #k cases (true_or_false (x0==x)) #H1 >H1 normalize // >(\P H1) >eqb_n_n % ] +qed. + +lemma vclose_vopen : ∀x,u.((νx.u)⌈x⌉) = u. #x #u @vclose_vopen_aux +qed. + +(* +theorem tm_to_tm : ∀t.tm' t → tm t. +#t #H elim H +*) + +lemma in_list_singleton : ∀T.∀t1,t2:T.t1 ∈ [t2] → t1 = t2. +#T #t1 #t2 #H @(in_list_inv_ind ??? H) /2/ +qed. + +lemma fresh_vclose_tm_aux : ∀u,x,k.x ∉ FV (vclose_tm_aux u x k). +#u #x elim u // +[ #n #k normalize cases (leb k n) normalize // +| #x0 #k normalize cases (true_or_false (x0==x)) #H >H normalize // + lapply (\Pf H) @not_to_not #Hin >(in_list_singleton ??? Hin) % +| #v1 #v2 #IH1 #IH2 #k normalize % #Hin cases (in_list_append_to_or_in_list ???? Hin) /2/ ] +qed. + +lemma fresh_vclose_tm : ∀u,x.x ∉ FV (νx.u). // +qed. + +lemma check_tm_true_to_toc : ∀u,k.check_tm u k = true → tm_or_ctx k u. +#u @(pretm_ind_plus ???? [ ] ? u) +[ #x #k #_ %2 +| #n #k change with (leb (S n) k) in ⊢ (??%?→?); #H % @leb_true_to_le // +| #v1 #v2 #rv1 #rv2 #k change with (pretm_ind_plus ???????) in ⊢ (??%?→?); + >pretm_ind_plus_app #H cases (andb_true ?? H) -H #Hv1 #Hv2 %3 + [ @rv1 @Hv1 | @rv2 @Hv2 ] +| #x #s #v #Hx #_ #rv #k change with (pretm_ind_plus ???????) in ⊢ (??%?→?); + >pretm_ind_plus_lam // #Hv %4 @rv @Hv ] +qed. + +lemma toc_to_check_tm_true : ∀u,k.tm_or_ctx k u → check_tm u k = true. +#u #k #Hu elim Hu // +[ #n #Hn change with (leb (S n) k) in ⊢ (??%?); @le_to_leb_true @Hn +| #v1 #v2 #Hv1 #Hv2 #IH1 #IH2 change with (pretm_ind_plus ???????) in ⊢ (??%?); + >pretm_ind_plus_app change with (check_tm v1 k ∧ check_tm v2 k) in ⊢ (??%?); /2/ +| #x #s #v #Hv #IH <(vclose_vopen x v) change with (pretm_ind_plus ???????) in ⊢ (??%?); + >pretm_ind_plus_lam [| // | @fresh_vclose_tm ] >(vclose_vopen x v) @IH ] +qed. + +lemma fresh_swap_tm : ∀z1,z2,u.z1 ∉ FV u → z2 ∉ FV u → swap 𝔸 z1 z2 · u = u. +#z1 #z2 #u elim u +[2: normalize in ⊢ (?→%→%→?); #x #Hz1 #Hz2 whd in ⊢ (??%?); >swap_other // + [ @(not_to_not … Hz2) | @(not_to_not … Hz1) ] // +|1: // +| #s #v #IH normalize #Hz1 #Hz2 >IH // [@Hz2|@Hz1] +| #v1 #v2 #IH1 #IH2 normalize #Hz1 #Hz2 + >IH1 [| @(not_to_not … Hz2) @in_list_to_in_list_append_l | @(not_to_not … Hz1) @in_list_to_in_list_append_l ] + >IH2 // [@(not_to_not … Hz2) @in_list_to_in_list_append_r | @(not_to_not … Hz1) @in_list_to_in_list_append_r ] +] +qed. + +theorem tm_to_tm2 : ∀u.tm u → tm2 u. +#t #Ht elim Ht +[ #n #Hn cases (not_le_Sn_O n) #Hfalse cases (Hfalse Hn) +| @tm_par +| #u #v #Hu #Hv @tm_app +| #x #s #u #Hu #IHu <(vclose_vopen x u) @tm_lam + [ @fresh_vclose_tm + | #y #Hy <(fresh_swap_tm x y (νx.u)) /2/ @fresh_vclose_tm ] +] +qed. + +theorem tm2_to_tm : ∀u.tm2 u → tm u. +#u #pu elim pu /2/ #x #s #v #Hx #Hv #IH %4 @IH // +qed. + +(* define PAR APP LAM *) +definition PAR ≝ λx.mk_TM (par x) ?. // qed. +definition APP ≝ λu,v:TM.mk_TM (app u v) ?. +change with (pretm_ind_plus ???????) in match (check_tm ??); >pretm_ind_plus_app +change with (check_tm ??) in match (pretm_ind_plus ???????); change with (check_tm ??) in match (pretm_ind_plus ???????) in ⊢ (??(??%)?); +@andb_elim >(tm_of_TM u) >(tm_of_TM v) % qed. +definition LAM ≝ λx,s.λu:TM.mk_TM (lam x s u) ?. +change with (pretm_ind_plus ???????) in match (check_tm ??); <(vclose_vopen x u) +>pretm_ind_plus_lam [| // | @fresh_vclose_tm ] +change with (check_tm ??) in match (pretm_ind_plus ???????); >vclose_vopen +@(tm_of_TM u) qed. + +axiom vopen_tm_down : ∀u,x,k.tm_or_ctx (S k) u → tm_or_ctx k (u⌈x⌉). +(* needs true_plus_false + +#u #x #k #Hu elim Hu +[ #n #Hn normalize cases (true_or_false (eqb n O)) #H >H [%2] + normalize >(?: leb n O = false) [|cases n in H; // >eqb_n_n #H destruct (H) ] + normalize lapply Hn cases n in H; normalize [ #Hfalse destruct (Hfalse) ] + #n0 #_ #Hn0 % @le_S_S_to_le // +| #x0 %2 +| #v1 #v2 #Hv1 #Hv2 #IH1 #IH2 %3 // +| #x0 #s #v #Hv #IH normalize @daemon +] +qed. +*) + +definition vopen_TM ≝ λu:CTX.λx.mk_TM (u⌈x⌉) ?. +@toc_to_check_tm_true @vopen_tm_down @check_tm_true_to_toc @ctx_of_CTX qed. + +axiom vclose_tm_up : ∀u,x,k.tm_or_ctx k u → tm_or_ctx (S k) (νx.u). + +definition vclose_TM ≝ λu:TM.λx.mk_CTX (νx.u) ?. +@toc_to_check_tm_true @vclose_tm_up @check_tm_true_to_toc @tm_of_TM qed. + +interpretation "ln wf term variable open" 'open u x = (vopen_TM u x). +interpretation "ln wf term variable close" 'nu x u = (vclose_TM u x). + +theorem tm_alpha : ∀x,y,s,u.x ∉ FV u → y ∉ FV u → lam x s (u⌈x⌉) = lam y s (u⌈y⌉). +#x #y #s #u #Hx #Hy whd in ⊢ (??%%); @eq_f >nominal_eta // >nominal_eta // +qed. + +lemma TM_to_tm2 : ∀u:TM.tm2 u. +#u @tm_to_tm2 @check_tm_true_to_toc @tm_of_TM qed. + +theorem TM_ind_plus_weak : + ∀P:pretm → Type[0]. + (∀x:𝔸.P (PAR x)) → + (∀v1,v2:TM.P v1 → P v2 → P (APP v1 v2)) → + ∀C:list 𝔸. + (∀x,s.∀v:CTX.x ∉ FV v → x ∉ C → + (∀y.y ∉ FV v → P (v⌈y⌉)) → P (LAM x s (v⌈x⌉))) → + ∀u:TM.P u. +#P #Hpar #Happ #C #Hlam #u elim (TM_to_tm2 u) // +[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 @(Happ (mk_TM …) (mk_TM …) IH1 IH2) + @toc_to_check_tm_true @tm2_to_tm // +| #x #s #v #Hx #pv #IH + lapply (p_fresh … (C@FV v)) letin x0 ≝ (N_fresh … (C@FV v)) #Hx0 + >(?:lam x s (v⌈x⌉) = lam x0 s (v⌈x0⌉)) + [|@tm_alpha // @(not_to_not … Hx0) @in_list_to_in_list_append_r ] + @(Hlam x0 s (mk_CTX v ?) ??) + [ <(nominal_eta … Hx) @toc_to_check_tm_true @vclose_tm_up @tm2_to_tm @pv // + | @(not_to_not … Hx0) @in_list_to_in_list_append_r + | @(not_to_not … Hx0) @in_list_to_in_list_append_l + | @IH ] +] +qed. + +lemma eq_mk_TM : ∀u,v.u = v → ∀pu,pv.mk_TM u pu = mk_TM v pv. +#u #v #Heq >Heq #pu #pv % +qed. + +lemma eq_P : ∀T:Type[0].∀t1,t2:T.t1 = t2 → ∀P:T → Type[0].P t1 → P t2. // qed. + +theorem TM_ind_plus : + ∀P:TM → Type[0]. + (∀x:𝔸.P (PAR x)) → + (∀v1,v2:TM.P v1 → P v2 → P (APP v1 v2)) → + ∀C:list 𝔸. + (∀x,s.∀v:CTX.x ∉ FV v → x ∉ C → + (∀y.y ∉ FV v → P (v⌈y⌉)) → P (LAM x s (v⌈x⌉))) → + ∀u:TM.P u. +#P #Hpar #Happ #C #Hlam * #u #pu +>(?:mk_TM u pu = + mk_TM u (toc_to_check_tm_true … (tm2_to_tm … (tm_to_tm2 … (check_tm_true_to_toc … pu))))) [|%] +elim (tm_to_tm2 u ?) // +[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 @(Happ (mk_TM …) (mk_TM …) IH1 IH2) +| #x #s #v #Hx #pv #IH + lapply (p_fresh … (C@FV v)) letin x0 ≝ (N_fresh … (C@FV v)) #Hx0 + lapply (Hlam x0 s (mk_CTX v ?) ???) + [2: @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_l @Hx0 + |4: @toc_to_check_tm_true <(nominal_eta x v) // @vclose_tm_up @tm2_to_tm @pv // + | #y #Hy whd in match (vopen_TM ??); + >(?:mk_TM (v⌈y⌉) ? = mk_TM (v⌈y⌉) (toc_to_check_tm_true (v⌈y⌉) O (tm2_to_tm (v⌈y⌉) (pv y Hy)))) + [@IH|%] + | @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_r @Hx0 + | @eq_P @eq_mk_TM whd in match (vopen_TM ??); @tm_alpha // @(not_to_not … Hx0) @in_list_to_in_list_append_r ] +] +qed. + +notation +"hvbox('nominal' u 'return' out 'with' + [ 'xpar' ident x ⇒ f1 + | 'xapp' ident v1 ident v2 ident recv1 ident recv2 ⇒ f2 + | 'xlam' ❨ident y # C❩ ident s ident w ident py1 ident py2 ident recw ⇒ f3 ])" +with precedence 48 +for @{ TM_ind_plus $out (λ${ident x}:?.$f1) + (λ${ident v1}:?.λ${ident v2}:?.λ${ident recv1}:?.λ${ident recv2}:?.$f2) + $C (λ${ident y}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?.λ${ident recw}:?.$f3) + $u }. + +(* include "basics/jmeq.ma".*) + +definition subst ≝ (λu:TM.λx,v. + nominal u return (λ_.TM) with + [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ PAR x0 ] (* u instead of PAR x0 does not work, u stays the same at every rec call! *) + | xapp v1 v2 recv1 recv2 ⇒ APP recv1 recv2 + | xlam ❨y # x::FV v❩ s w py1 py2 recw ⇒ LAM y s (recw y py1) ]). + +lemma subst_def : ∀u,x,v.subst u x v = + nominal u return (λ_.TM) with + [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ PAR x0 ] + | xapp v1 v2 recv1 recv2 ⇒ APP recv1 recv2 + | xlam ❨y # x::FV v❩ s w py1 py2 recw ⇒ LAM y s (recw y py1) ]. // +qed. + +axiom TM_ind_plus_LAM : + ∀x,s,u,out,f1,f2,C,f3,Hx1,Hx2. + TM_ind_plus out f1 f2 C f3 (LAM x s (u⌈x⌉)) = + f3 x s u Hx1 Hx2 (λy,Hy.TM_ind_plus ? f1 f2 C f3 ?). + +axiom TM_ind_plus_APP : + ∀u1,u2,out,f1,f2,C,f3. + TM_ind_plus out f1 f2 C f3 (APP u1 u2) = + f2 u1 u2 (TM_ind_plus out f1 f2 C f3 ?) (TM_ind_plus out f1 f2 C f3 ?). + +lemma eq_mk_CTX : ∀u,v.u = v → ∀pu,pv.mk_CTX u pu = mk_CTX v pv. +#u #v #Heq >Heq #pu #pv % +qed. + +lemma vclose_vopen_TM : ∀x.∀u:TM.((νx.u)⌈x⌉) = u. +#x * #u #pu @eq_mk_TM @vclose_vopen qed. + +lemma nominal_eta_CTX : ∀x.∀u:CTX.x ∉ FV u → (νx.(u⌈x⌉)) = u. +#x * #u #pu #Hx @eq_mk_CTX @nominal_eta // qed. + +theorem TM_alpha : ∀x,y,s.∀u:CTX.x ∉ FV u → y ∉ FV u → LAM x s (u⌈x⌉) = LAM y s (u⌈y⌉). +#x #y #s #u #Hx #Hy @eq_mk_TM @tm_alpha // qed. + +axiom in_vopen_CTX : ∀x,y.∀v:CTX.x ∈ FV (v⌈y⌉) → x = y ∨ x ∈ FV v. + +theorem subst_fresh : ∀u,v:TM.∀x.x ∉ FV u → subst u x v = u. +#u #v #x @(TM_ind_plus … (x::FV v) … u) +[ #x0 normalize in ⊢ (%→?); #Hx normalize in ⊢ (??%?); + >(\bf ?) [| @(not_to_not … Hx) #Heq >Heq % ] % +| #u1 #u2 #IH1 #IH2 normalize in ⊢ (%→?); #Hx + >subst_def >TM_ind_plus_APP @eq_mk_TM @eq_f2 @eq_f + [ subst_def >TM_ind_plus_LAM [|@HC|@Hx0] + @eq_f nominal_eta_CTX // + cases (in_vopen_CTX … Hx) // #Heq >Heq in HC; * #HC @False_ind @HC % +] +qed. + +example subst_LAM_same : ∀x,s,u,v. subst (LAM x s u) x v = LAM x s u. +#x #s #u #v >subst_def <(vclose_vopen_TM x u) +lapply (p_fresh … (FV (νx.u)@x::FV v)) letin x0 ≝ (N_fresh … (FV (νx.u)@x::FV v)) #Hx0 +>(TM_alpha x x0) +[| @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_l @Hx0 | @fresh_vclose_tm ] +>TM_ind_plus_LAM [| @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_r @Hx0 | @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_l @Hx0 ] +@eq_f change with (subst ((νx.u)⌈x0⌉) x v) in ⊢ (??%?); @subst_fresh +@(not_to_not … Hx0) #Hx0' cases (in_vopen_CTX … Hx0') +[ #Heq >Heq @in_list_to_in_list_append_r % +| #Hfalse @False_ind cases (fresh_vclose_tm u x) #H @H @Hfalse ] +qed. + +(* +notation > "Λ ident x. ident T [ident x] ↦ P" + with precedence 48 for @{'foo (λ${ident x}.λ${ident T}.$P)}. + +notation < "Λ ident x. ident T [ident x] ↦ P" + with precedence 48 for @{'foo (λ${ident x}:$Q.λ${ident T}:$R.$P)}. +*) + +(* +notation +"hvbox('nominal' u 'with' + [ 'xpar' ident x ⇒ f1 + | 'xapp' ident v1 ident v2 ⇒ f2 + | 'xlam' ident x # C s w ⇒ f3 ])" +with precedence 48 +for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1) + (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) + $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. +*) + +(* +notation +"hvbox('nominal' u 'with' + [ 'xpar' ident x ^ f1 + | 'xapp' ident v1 ident v2 ^ f2 ])" +(* | 'xlam' ident x # C s w ^ f3 ]) *) +with precedence 48 +for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1) + (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) + $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. +*) +notation +"hvbox('nominal' u 'with' + [ 'xpar' ident x ^ f1 + | 'xapp' ident v1 ident v2 ^ f2 ])" +with precedence 48 +for @{ tm2_ind_plus ? (λ${ident x}:?.$f1) + (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) + $C (λ${ident x}:?.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. + +axiom in_Env : 𝔸 × tp → Env → Prop. +notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}. +interpretation "Env membership" 'lefttriangle x l = (in_Env x l). + + + +inductive judg : list tp → tm → tp → Prop ≝ +| t_var : ∀g,n,t.Nth ? n g = Some ? t → judg g (var n) t +| t_app : ∀g,m,n,t,u.judg g m (arr t u) → judg g n t → judg g (app m n) u +| t_abs : ∀g,t,m,u.judg (t::g) m u → judg g (abs t m) (arr t u). + +definition Env := list (𝔸 × tp). + +axiom vclose_env : Env → list tp. +axiom vclose_tm : Env → tm → tm. +axiom Lam : 𝔸 → tp → tm → tm. +definition Judg ≝ λG,M,T.judg (vclose_env G) (vclose_tm G M) T. +definition dom ≝ λG:Env.map ?? (fst ??) G. + +definition sctx ≝ 𝔸 × tm. +axiom swap_tm : 𝔸 → 𝔸 → tm → tm. +definition sctx_app : sctx → 𝔸 → tm ≝ λM0,Y.let 〈X,M〉 ≝ M0 in swap_tm X Y M. + +axiom in_list : ∀A:Type[0].A → list A → Prop. +interpretation "list membership" 'mem x l = (in_list ? x l). +interpretation "list non-membership" 'notmem x l = (Not (in_list ? x l)). + +axiom in_Env : 𝔸 × tp → Env → Prop. +notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}. +interpretation "Env membership" 'lefttriangle x l = (in_Env x l). + +(* axiom Lookup : 𝔸 → Env → option tp. *) + +(* forma alto livello del judgment + t_abs* : ∀G,T,X,M,U. + (∀Y ∉ supp(M).Judg (〈Y,T〉::G) (M[Y]) U) → + Judg G (Lam X T (M[X])) (arr T U) *) + +(* prima dimostrare, poi perfezionare gli assiomi, poi dimostrarli *) + +axiom Judg_ind : ∀P:Env → tm → tp → Prop. + (∀X,G,T.〈X,T〉 ◃ G → P G (par X) T) → + (∀G,M,N,T,U. + Judg G M (arr T U) → Judg G N T → + P G M (arr T U) → P G N T → P G (app M N) U) → + (∀G,T1,T2,X,M1. + (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → Judg (〈Y,T1〉::G) (sctx_app M1 Y) T2) → + (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → P (〈Y,T1〉::G) (sctx_app M1 Y) T2) → + P G (Lam X T1 (sctx_app M1 X)) (arr T1 T2)) → + ∀G,M,T.Judg G M T → P G M T. + +axiom t_par : ∀X,G,T.〈X,T〉 ◃ G → Judg G (par X) T. +axiom t_app2 : ∀G,M,N,T,U.Judg G M (arr T U) → Judg G N T → Judg G (app M N) U. +axiom t_Lam : ∀G,X,M,T,U.Judg (〈X,T〉::G) M U → Judg G (Lam X T M) (arr T U). + +definition subenv ≝ λG1,G2.∀x.x ◃ G1 → x ◃ G2. +interpretation "subenv" 'subseteq G1 G2 = (subenv G1 G2). + +axiom daemon : ∀P:Prop.P. + +theorem weakening : ∀G1,G2,M,T.G1 ⊆ G2 → Judg G1 M T → Judg G2 M T. +#G1 #G2 #M #T #Hsub #HJ lapply Hsub lapply G2 -G2 change with (∀G2.?) +@(Judg_ind … HJ) +[ #X #G #T0 #Hin #G2 #Hsub @t_par @Hsub // +| #G #M0 #N #T0 #U #HM0 #HN #IH1 #IH2 #G2 #Hsub @t_app2 + [| @IH1 // | @IH2 // ] +| #G #T1 #T2 #X #M1 #HM1 #IH #G2 #Hsub @t_Lam @IH + [ (* trivial property of Lam *) @daemon + | (* trivial property of subenv *) @daemon ] +] +qed. + +(* Serve un tipo Tm per i termini localmente chiusi e i suoi principi di induzione e + ricorsione *) \ No newline at end of file diff --git a/matita/matita/broken_lib/binding/ln_concrete.ma b/matita/matita/broken_lib/binding/ln_concrete.ma new file mode 100644 index 000000000..bdafdb1a0 --- /dev/null +++ b/matita/matita/broken_lib/binding/ln_concrete.ma @@ -0,0 +1,683 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basics/lists/list.ma". +include "basics/deqsets.ma". +include "binding/names.ma". +include "binding/fp.ma". + +definition alpha : Nset ≝ X. check alpha +notation "𝔸" non associative with precedence 90 for @{'alphabet}. +interpretation "set of names" 'alphabet = alpha. + +inductive tp : Type[0] ≝ +| top : tp +| arr : tp → tp → tp. +inductive pretm : Type[0] ≝ +| var : nat → pretm +| par : 𝔸 → pretm +| abs : tp → pretm → pretm +| app : pretm → pretm → pretm. + +let rec Nth T n (l:list T) on n ≝ + match l with + [ nil ⇒ None ? + | cons hd tl ⇒ match n with + [ O ⇒ Some ? hd + | S n0 ⇒ Nth T n0 tl ] ]. + +let rec vclose_tm_aux u x k ≝ match u with + [ var n ⇒ if (leb k n) then var (S n) else u + | par x0 ⇒ if (x0 == x) then (var k) else u + | app v1 v2 ⇒ app (vclose_tm_aux v1 x k) (vclose_tm_aux v2 x k) + | abs s v ⇒ abs s (vclose_tm_aux v x (S k)) ]. +definition vclose_tm ≝ λu,x.vclose_tm_aux u x O. + +definition vopen_var ≝ λn,x,k.match eqb n k with + [ true ⇒ par x + | false ⇒ match leb n k with + [ true ⇒ var n + | false ⇒ var (pred n) ] ]. + +let rec vopen_tm_aux u x k ≝ match u with + [ var n ⇒ vopen_var n x k + | par x0 ⇒ u + | app v1 v2 ⇒ app (vopen_tm_aux v1 x k) (vopen_tm_aux v2 x k) + | abs s v ⇒ abs s (vopen_tm_aux v x (S k)) ]. +definition vopen_tm ≝ λu,x.vopen_tm_aux u x O. + +let rec FV u ≝ match u with + [ par x ⇒ [x] + | app v1 v2 ⇒ FV v1@FV v2 + | abs s v ⇒ FV v + | _ ⇒ [ ] ]. + +definition lam ≝ λx,s,u.abs s (vclose_tm u x). + +let rec Pi_map_tm p u on u ≝ match u with +[ par x ⇒ par (p x) +| var _ ⇒ u +| app v1 v2 ⇒ app (Pi_map_tm p v1) (Pi_map_tm p v2) +| abs s v ⇒ abs s (Pi_map_tm p v) ]. + +interpretation "permutation of tm" 'middot p x = (Pi_map_tm p x). + +notation "hvbox(u⌈x⌉)" + with precedence 45 + for @{ 'open $u $x }. + +(* +notation "hvbox(u⌈x⌉)" + with precedence 45 + for @{ 'open $u $x }. +notation "❴ u ❵ x" non associative with precedence 90 for @{ 'open $u $x }. +*) +interpretation "ln term variable open" 'open u x = (vopen_tm u x). +notation < "hvbox(ν x break . u)" + with precedence 20 +for @{'nu $x $u }. +notation > "ν list1 x sep , . term 19 u" with precedence 20 + for ${ fold right @{$u} rec acc @{'nu $x $acc)} }. +interpretation "ln term variable close" 'nu x u = (vclose_tm u x). + +let rec tm_height u ≝ match u with +[ app v1 v2 ⇒ S (max (tm_height v1) (tm_height v2)) +| abs s v ⇒ S (tm_height v) +| _ ⇒ O ]. + +theorem le_n_O_rect_Type0 : ∀n:nat. n ≤ O → ∀P: nat →Type[0]. P O → P n. +#n (cases n) // #a #abs cases (?:False) /2/ qed. + +theorem nat_rect_Type0_1 : ∀n:nat.∀P:nat → Type[0]. +(∀m.(∀p. p < m → P p) → P m) → P n. +#n #P #H +cut (∀q:nat. q ≤ n → P q) /2/ +(elim n) + [#q #HleO (* applica male *) + @(le_n_O_rect_Type0 ? HleO) + @H #p #ltpO cases (?:False) /2/ (* 3 *) + |#p #Hind #q #HleS + @H #a #lta @Hind @le_S_S_to_le /2/ + ] +qed. + +lemma leb_false_to_lt : ∀n,m. leb n m = false → m < n. +#n elim n +[ #m normalize #H destruct(H) +| #n0 #IH * // #m normalize #H @le_S_S @IH // ] +qed. + +lemma nominal_eta_aux : ∀x,u.x ∉ FV u → ∀k.vclose_tm_aux (vopen_tm_aux u x k) x k = u. +#x #u elim u +[ #n #_ #k normalize cases (decidable_eq_nat n k) #Hnk + [ >Hnk >eqb_n_n whd in ⊢ (??%?); >(\b ?) // + | >(not_eq_to_eqb_false … Hnk) normalize cases (true_or_false (leb n k)) #Hleb + [ >Hleb normalize >(?:leb k n = false) // + @lt_to_leb_false @not_eq_to_le_to_lt /2/ + | >Hleb normalize >(?:leb k (pred n) = true) normalize + [ cases (leb_false_to_lt … Hleb) // + | @le_to_leb_true cases (leb_false_to_lt … Hleb) normalize /2/ ] ] ] +| #y normalize in ⊢ (%→?→?); #Hy whd in ⊢ (?→??%?); >(\bf ?) // @(not_to_not … Hy) // +| #s #v #IH normalize #Hv #k >IH // @Hv +| #v1 #v2 #IH1 #IH2 normalize #Hv1v2 #k + >IH1 [ >IH2 // | @(not_to_not … Hv1v2) @in_list_to_in_list_append_l ] + @(not_to_not … Hv1v2) @in_list_to_in_list_append_r ] +qed. + +corollary nominal_eta : ∀x,u.x ∉ FV u → (νx.u⌈x⌉) = u. +#x #u #Hu @nominal_eta_aux // +qed. + +lemma eq_height_vopen_aux : ∀v,x,k.tm_height (vopen_tm_aux v x k) = tm_height v. +#v #x elim v +[ #n #k normalize cases (eqb n k) // cases (leb n k) // +| #u #k % +| #s #u #IH #k normalize >IH % +| #u1 #u2 #IH1 #IH2 #k normalize >IH1 >IH2 % ] +qed. + +corollary eq_height_vopen : ∀v,x.tm_height (v⌈x⌉) = tm_height v. +#v #x @eq_height_vopen_aux +qed. + +theorem pretm_ind_plus_aux : + ∀P:pretm → Type[0]. + (∀x:𝔸.P (par x)) → + (∀n:ℕ.P (var n)) → + (∀v1,v2. P v1 → P v2 → P (app v1 v2)) → + ∀C:list 𝔸. + (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) → + ∀n,u.tm_height u ≤ n → P u. +#P #Hpar #Hvar #Happ #C #Hlam #n change with ((λn.?) n); @(nat_rect_Type0_1 n ??) +#m cases m +[ #_ * /2/ + [ normalize #s #v #Hfalse cases (?:False) cases (not_le_Sn_O (tm_height v)) /2/ + | #v1 #v2 whd in ⊢ (?%?→?); #Hfalse cases (?:False) cases (not_le_Sn_O (S (max ??))) /2/ ] ] +-m #m #IH * /2/ +[ #s #v whd in ⊢ (?%?→?); #Hv + lapply (p_fresh … (C@FV v)) letin y ≝ (N_fresh … (C@FV v)) #Hy + >(?:abs s v = lam y s (v⌈y⌉)) + [| whd in ⊢ (???%); >nominal_eta // @(not_to_not … Hy) @in_list_to_in_list_append_r ] + @Hlam + [ @(not_to_not … Hy) @in_list_to_in_list_append_r + | @(not_to_not … Hy) @in_list_to_in_list_append_l ] + @IH [| @Hv | >eq_height_vopen % ] +| #v1 #v2 whd in ⊢ (?%?→?); #Hv @Happ + [ @IH [| @Hv | // ] | @IH [| @Hv | // ] ] ] +qed. + +corollary pretm_ind_plus : + ∀P:pretm → Type[0]. + (∀x:𝔸.P (par x)) → + (∀n:ℕ.P (var n)) → + (∀v1,v2. P v1 → P v2 → P (app v1 v2)) → + ∀C:list 𝔸. + (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) → + ∀u.P u. +#P #Hpar #Hvar #Happ #C #Hlam #u @pretm_ind_plus_aux /2/ +qed. + +(* maps a permutation to a list of terms *) +definition Pi_map_list : (𝔸 → 𝔸) → list 𝔸 → list 𝔸 ≝ map 𝔸 𝔸 . + +(* interpretation "permutation of name list" 'middot p x = (Pi_map_list p x).*) + +(* +inductive tm : pretm → Prop ≝ +| tm_par : ∀x:𝔸.tm (par x) +| tm_app : ∀u,v.tm u → tm v → tm (app u v) +| tm_lam : ∀x,s,u.tm u → tm (lam x s u). + +inductive ctx_aux : nat → pretm → Prop ≝ +| ctx_var : ∀n,k.n < k → ctx_aux k (var n) +| ctx_par : ∀x,k.ctx_aux k (par x) +| ctx_app : ∀u,v,k.ctx_aux k u → ctx_aux k v → ctx_aux k (app u v) +(* è sostituibile da ctx_lam ? *) +| ctx_abs : ∀s,u.ctx_aux (S k) u → ctx_aux k (abs s u). +*) + +inductive tm_or_ctx (k:nat) : pretm → Type[0] ≝ +| toc_var : ∀n.n < k → tm_or_ctx k (var n) +| toc_par : ∀x.tm_or_ctx k (par x) +| toc_app : ∀u,v.tm_or_ctx k u → tm_or_ctx k v → tm_or_ctx k (app u v) +| toc_lam : ∀x,s,u.tm_or_ctx k u → tm_or_ctx k (lam x s u). + +definition tm ≝ λt.tm_or_ctx O t. +definition ctx ≝ λt.tm_or_ctx 1 t. + +record TM : Type[0] ≝ { + pretm_of_TM :> pretm; + tm_of_TM : tm pretm_of_TM +}. + +record CTX : Type[0] ≝ { + pretm_of_CTX :> pretm; + ctx_of_CTX : ctx pretm_of_CTX +}. + +inductive tm2 : pretm → Type[0] ≝ +| tm_par : ∀x.tm2 (par x) +| tm_app : ∀u,v.tm2 u → tm2 v → tm2 (app u v) +| tm_lam : ∀x,s,u.x ∉ FV u → (∀y.y ∉ FV u → tm2 (u⌈y⌉)) → tm2 (lam x s (u⌈x⌉)). + +(* +inductive tm' : pretm → Prop ≝ +| tm_par : ∀x.tm' (par x) +| tm_app : ∀u,v.tm' u → tm' v → tm' (app u v) +| tm_lam : ∀x,s,u,C.x ∉ FV u → x ∉ C → (∀y.y ∉ FV u → tm' (❴u❵y)) → tm' (lam x s (❴u❵x)). +*) + +axiom swap_inj : ∀N.∀z1,z2,x,y.swap N z1 z2 x = swap N z1 z2 y → x = y. + +lemma pi_vclose_tm : + ∀z1,z2,x,u.swap 𝔸 z1 z2·(νx.u) = (ν swap ? z1 z2 x.swap 𝔸 z1 z2 · u). +#z1 #z2 #x #u +change with (vclose_tm_aux ???) in match (vclose_tm ??); +change with (vclose_tm_aux ???) in ⊢ (???%); lapply O elim u +[3:whd in ⊢ (?→?→(?→ ??%%)→?→??%%); // +|4:whd in ⊢ (?→?→(?→??%%)→(?→??%%)→?→??%%); // +| #n #k whd in ⊢ (??(??%)%); cases (leb k n) normalize % +| #x0 #k cases (true_or_false (x0==z1)) #H1 >H1 whd in ⊢ (??%%); + [ cases (true_or_false (x0==x)) #H2 >H2 whd in ⊢ (??(??%)%); + [ <(\P H2) >H1 whd in ⊢ (??(??%)%); >(\b ?) // >(\b ?) // + | >H2 whd in match (swap ????); >H1 + whd in match (if false then var k else ?); + whd in match (if true then z2 else ?); >(\bf ?) + [ >(\P H1) >swap_left % + | <(swap_inv ? z1 z2 z2) in ⊢ (?(??%?)); % #H3 + lapply (swap_inj … H3) >swap_right #H4

H1 #H destruct (H) ] + ] + | >(?:(swap ? z1 z2 x0 == swap ? z1 z2 x) = (x0 == x)) + [| cases (true_or_false (x0==x)) #H2 >H2 + [ >(\P H2) @(\b ?) % + | @(\bf ?) % #H >(swap_inj … H) in H2; >(\b ?) // #H0 destruct (H0) ] ] + cases (true_or_false (x0==x)) #H2 >H2 whd in ⊢ (??(??%)%); + [ <(\P H2) >H1 >(\b (refl ??)) % + | >H1 >H2 % ] + ] + ] +qed. + +lemma pi_vopen_tm : + ∀z1,z2,x,u.swap 𝔸 z1 z2·(u⌈x⌉) = (swap 𝔸 z1 z2 · u⌈swap 𝔸 z1 z2 x⌉). +#z1 #z2 #x #u +change with (vopen_tm_aux ???) in match (vopen_tm ??); +change with (vopen_tm_aux ???) in ⊢ (???%); lapply O elim u // +[2: #s #v whd in ⊢ ((?→??%%)→?→??%%); // +|3: #v1 #v2 whd in ⊢ ((?→??%%)→(?→??%%)→?→??%%); /2/ ] +#n #k whd in ⊢ (??(??%)%); cases (true_or_false (eqb n k)) #H1 >H1 // +cases (true_or_false (leb n k)) #H2 >H2 normalize // +qed. + +lemma pi_lam : + ∀z1,z2,x,s,u.swap 𝔸 z1 z2 · lam x s u = lam (swap 𝔸 z1 z2 x) s (swap 𝔸 z1 z2 · u). +#z1 #z2 #x #s #u whd in ⊢ (???%); <(pi_vclose_tm …) % +qed. + +lemma eqv_FV : ∀z1,z2,u.FV (swap 𝔸 z1 z2 · u) = Pi_map_list (swap 𝔸 z1 z2) (FV u). +#z1 #z2 #u elim u // +[ #s #v #H @H +| #v1 #v2 whd in ⊢ (??%%→??%%→??%%); #H1 #H2 >H1 >H2 + whd in ⊢ (???(????%)); /2/ ] +qed. + +lemma swap_inv_tm : ∀z1,z2,u.swap 𝔸 z1 z2 · (swap 𝔸 z1 z2 · u) = u. +#z1 #z2 #u elim u +[1: #n % +|3: #s #v whd in ⊢ (?→??%%); // +|4: #v1 #v2 #Hv1 #Hv2 whd in ⊢ (??%%); // ] +#x whd in ⊢ (??%?); >swap_inv % +qed. + +lemma eqv_in_list : ∀x,l,z1,z2.x ∈ l → swap 𝔸 z1 z2 x ∈ Pi_map_list (swap 𝔸 z1 z2) l. +#x #l #z1 #z2 #Hin elim Hin +[ #x0 #l0 % +| #x1 #x2 #l0 #Hin #IH %2 @IH ] +qed. + +lemma eqv_tm2 : ∀u.tm2 u → ∀z1,z2.tm2 ((swap ? z1 z2)·u). +#u #Hu #z1 #z2 letin p ≝ (swap ? z1 z2) elim Hu /2/ +#x #s #v #Hx #Hv #IH >pi_lam >pi_vopen_tm %3 +[ @(not_to_not … Hx) -Hx #Hx + <(swap_inv ? z1 z2 x) <(swap_inv_tm z1 z2 v) >eqv_FV @eqv_in_list // +| #y #Hy <(swap_inv ? z1 z2 y) + eqv_FV @eqv_in_list // +] +qed. + +lemma vclose_vopen_aux : ∀x,u,k.vopen_tm_aux (vclose_tm_aux u x k) x k = u. +#x #u elim u [1,3,4:normalize //] +[ #n #k cases (true_or_false (leb k n)) #H >H whd in ⊢ (??%?); + [ cases (true_or_false (eqb (S n) k)) #H1 >H1 + [ <(eqb_true_to_eq … H1) in H; #H lapply (leb_true_to_le … H) -H #H + cases (le_to_not_lt … H) -H #H cases (H ?) % + | whd in ⊢ (??%?); >lt_to_leb_false // @le_S_S /2/ ] + | cases (true_or_false (eqb n k)) #H1 >H1 normalize + [ >(eqb_true_to_eq … H1) in H; #H lapply (leb_false_to_not_le … H) -H + * #H cases (H ?) % + | >le_to_leb_true // @not_lt_to_le % #H2 >le_to_leb_true in H; + [ #H destruct (H) | /2/ ] + ] + ] +| #x0 #k whd in ⊢ (??(?%??)?); cases (true_or_false (x0==x)) + #H1 >H1 normalize // >(\P H1) >eqb_n_n % ] +qed. + +lemma vclose_vopen : ∀x,u.((νx.u)⌈x⌉) = u. #x #u @vclose_vopen_aux +qed. + +(* +theorem tm_to_tm : ∀t.tm' t → tm t. +#t #H elim H +*) + +lemma in_list_singleton : ∀T.∀t1,t2:T.t1 ∈ [t2] → t1 = t2. +#T #t1 #t2 #H @(in_list_inv_ind ??? H) /2/ +qed. + +lemma fresh_vclose_tm_aux : ∀u,x,k.x ∉ FV (vclose_tm_aux u x k). +#u #x elim u // +[ #n #k normalize cases (leb k n) normalize // +| #x0 #k whd in ⊢ (?(???(?%))); cases (true_or_false (x0==x)) #H >H normalize // + lapply (\Pf H) @not_to_not #Hin >(in_list_singleton ??? Hin) % +| #v1 #v2 #IH1 #IH2 #k normalize % #Hin cases (in_list_append_to_or_in_list ???? Hin) -Hin #Hin + [ cases (IH1 k) -IH1 #IH1 @IH1 @Hin | cases (IH2 k) -IH2 #IH2 @IH2 @Hin ] +qed. + +lemma fresh_vclose_tm : ∀u,x.x ∉ FV (νx.u). // +qed. + +lemma fresh_swap_tm : ∀z1,z2,u.z1 ∉ FV u → z2 ∉ FV u → swap 𝔸 z1 z2 · u = u. +#z1 #z2 #u elim u +[2: normalize in ⊢ (?→%→%→?); #x #Hz1 #Hz2 whd in ⊢ (??%?); >swap_other // + [ @(not_to_not … Hz2) | @(not_to_not … Hz1) ] // +|1: // +| #s #v #IH normalize #Hz1 #Hz2 >IH // [@Hz2|@Hz1] +| #v1 #v2 #IH1 #IH2 normalize #Hz1 #Hz2 + >IH1 [| @(not_to_not … Hz2) @in_list_to_in_list_append_l | @(not_to_not … Hz1) @in_list_to_in_list_append_l ] + >IH2 // [@(not_to_not … Hz2) @in_list_to_in_list_append_r | @(not_to_not … Hz1) @in_list_to_in_list_append_r ] +] +qed. + +theorem tm_to_tm2 : ∀u.tm u → tm2 u. +#t #Ht elim Ht +[ #n #Hn cases (not_le_Sn_O n) #Hfalse cases (Hfalse Hn) +| @tm_par +| #u #v #Hu #Hv @tm_app +| #x #s #u #Hu #IHu <(vclose_vopen x u) @tm_lam + [ @fresh_vclose_tm + | #y #Hy <(fresh_swap_tm x y (νx.u)) /2/ @fresh_vclose_tm ] +] +qed. + +theorem tm2_to_tm : ∀u.tm2 u → tm u. +#u #pu elim pu /2/ #x #s #v #Hx #Hv #IH %4 @IH // +qed. + +definition PAR ≝ λx.mk_TM (par x) ?. // qed. +definition APP ≝ λu,v:TM.mk_TM (app u v) ?./2/ qed. +definition LAM ≝ λx,s.λu:TM.mk_TM (lam x s u) ?./2/ qed. + +axiom vopen_tm_down : ∀u,x,k.tm_or_ctx (S k) u → tm_or_ctx k (u⌈x⌉). +(* needs true_plus_false + +#u #x #k #Hu elim Hu +[ #n #Hn normalize cases (true_or_false (eqb n O)) #H >H [%2] + normalize >(?: leb n O = false) [|cases n in H; // >eqb_n_n #H destruct (H) ] + normalize lapply Hn cases n in H; normalize [ #Hfalse destruct (Hfalse) ] + #n0 #_ #Hn0 % @le_S_S_to_le // +| #x0 %2 +| #v1 #v2 #Hv1 #Hv2 #IH1 #IH2 %3 // +| #x0 #s #v #Hv #IH normalize @daemon +] +qed. +*) + +definition vopen_TM ≝ λu:CTX.λx.mk_TM (u⌈x⌉) (vopen_tm_down …). @ctx_of_CTX qed. + +axiom vclose_tm_up : ∀u,x,k.tm_or_ctx k u → tm_or_ctx (S k) (νx.u). + +definition vclose_TM ≝ λu:TM.λx.mk_CTX (νx.u) (vclose_tm_up …). @tm_of_TM qed. + +interpretation "ln wf term variable open" 'open u x = (vopen_TM u x). +interpretation "ln wf term variable close" 'nu x u = (vclose_TM u x). + +theorem tm_alpha : ∀x,y,s,u.x ∉ FV u → y ∉ FV u → lam x s (u⌈x⌉) = lam y s (u⌈y⌉). +#x #y #s #u #Hx #Hy whd in ⊢ (??%%); @eq_f >nominal_eta // >nominal_eta // +qed. + +theorem TM_ind_plus : +(* non si può dare il principio in modo dipendente (almeno utilizzando tm2) + la "prova" purtroppo è in Type e non si può garantire che sia esattamente + quella che ci aspetteremmo + *) + ∀P:pretm → Type[0]. + (∀x:𝔸.P (PAR x)) → + (∀v1,v2:TM.P v1 → P v2 → P (APP v1 v2)) → + ∀C:list 𝔸. + (∀x,s.∀v:CTX.x ∉ FV v → x ∉ C → + (∀y.y ∉ FV v → P (v⌈y⌉)) → P (LAM x s (v⌈x⌉))) → + ∀u:TM.P u. +#P #Hpar #Happ #C #Hlam * #u #pu elim (tm_to_tm2 u pu) // +[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 @(Happ (mk_TM …) (mk_TM …)) /2/ +| #x #s #v #Hx #pv #IH + lapply (p_fresh … (C@FV v)) letin x0 ≝ (N_fresh … (C@FV v)) #Hx0 + >(?:lam x s (v⌈x⌉) = lam x0 s (v⌈x0⌉)) + [|@tm_alpha // @(not_to_not … Hx0) @in_list_to_in_list_append_r ] + @(Hlam x0 s (mk_CTX v ?) ??) + [ <(nominal_eta … Hx) @vclose_tm_up @tm2_to_tm @pv // + | @(not_to_not … Hx0) @in_list_to_in_list_append_r + | @(not_to_not … Hx0) @in_list_to_in_list_append_l + | @IH ] +] +qed. + +notation +"hvbox('nominal' u 'return' out 'with' + [ 'xpar' ident x ⇒ f1 + | 'xapp' ident v1 ident v2 ident recv1 ident recv2 ⇒ f2 + | 'xlam' ❨ident y # C❩ ident s ident w ident py1 ident py2 ident recw ⇒ f3 ])" +with precedence 48 +for @{ TM_ind_plus $out (λ${ident x}:?.$f1) + (λ${ident v1}:?.λ${ident v2}:?.λ${ident recv1}:?.λ${ident recv2}:?.$f2) + $C (λ${ident y}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?.λ${ident recw}:?.$f3) + $u }. + +(* include "basics/jmeq.ma".*) + +definition subst ≝ (λu:TM.λx,v. + nominal u return (λ_.TM) with + [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ u ] + | xapp v1 v2 recv1 recv2 ⇒ APP recv1 recv2 + | xlam ❨y # x::FV v❩ s w py1 py2 recw ⇒ LAM y s (recw y py1) ]). + +lemma fasfd : ∀s,v. pretm_of_TM (subst (LAM O s (PAR 1)) O v) = pretm_of_TM (LAM O s (PAR 1)). +#s #v normalize in ⊢ (??%?); + + +theorem tm2_ind_plus : +(* non si può dare il principio in modo dipendente (almeno utilizzando tm2) *) + ∀P:pretm → Type[0]. + (∀x:𝔸.P (par x)) → + (∀v1,v2.tm2 v1 → tm2 v2 → P v1 → P v2 → P (app v1 v2)) → + ∀C:list 𝔸. + (∀x,s,v.x ∉ FV v → x ∉ C → (∀y.y ∉ FV v → tm2 (v⌈y⌉)) → + (∀y.y ∉ FV v → P (v⌈y⌉)) → P (lam x s (v⌈x⌉))) → + ∀u.tm2 u → P u. +#P #Hpar #Happ #C #Hlam #u #pu elim pu /2/ +#x #s #v #px #pv #IH +lapply (p_fresh … (C@FV v)) letin y ≝ (N_fresh … (C@FV v)) #Hy +>(?:lam x s (v⌈x⌉) = lam y s (v⌈y⌉)) [| @tm_alpha // @(not_to_not … Hy) @in_list_to_in_list_append_r ] +@Hlam /2/ lapply Hy -Hy @not_to_not #Hy +[ @in_list_to_in_list_append_r @Hy | @in_list_to_in_list_append_l @Hy ] +qed. + +definition check_tm ≝ + λu.pretm_ind_plus ? (λ_.true) (λ_.false) + (λv1,v2,r1,r2.r1 ∧ r2) [ ] (λx,s,v,pv1,pv2,rv.rv) u. + +(* +lemma check_tm_complete : ∀u.tm u → check_tm u = true. +#u #pu @(tm2_ind_plus … [ ] … (tm_to_tm2 ? pu)) // +[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 +| #x #s #v #Hx1 #Hx2 #Hv #IH +*) + +notation +"hvbox('nominal' u 'return' out 'with' + [ 'xpar' ident x ⇒ f1 + | 'xapp' ident v1 ident v2 ident pv1 ident pv2 ident recv1 ident recv2 ⇒ f2 + | 'xlam' ❨ident y # C❩ ident s ident w ident py1 ident py2 ident pw ident recw ⇒ f3 ])" +with precedence 48 +for @{ tm2_ind_plus $out (λ${ident x}:?.$f1) + (λ${ident v1}:?.λ${ident v2}:?.λ${ident pv1}:?.λ${ident pv2}:?.λ${ident recv1}:?.λ${ident recv2}:?.$f2) + $C (λ${ident y}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?.λ${ident pw}:?.λ${ident recw}:?.$f3) + ? (tm_to_tm2 ? $u) }. +(* notation +"hvbox('nominal' u 'with' + [ 'xlam' ident x # C ident s ident w ⇒ f3 ])" +with precedence 48 +for @{ tm2_ind_plus ??? + $C (λ${ident x}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?. + λ${ident pw}:?.λ${ident recw}:?.$f3) $u (tm_to_tm2 ??) }. +*) + + +definition subst ≝ (λu.λpu:tm u.λx,v. + nominal pu return (λ_.pretm) with + [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ u ] + | xapp v1 v2 pv1 pv2 recv1 recv2 ⇒ app recv1 recv2 + | xlam ❨y # x::FV v❩ s w py1 py2 pw recw ⇒ lam y s (recw y py1) ]). + +lemma fasfd : ∀x,s,u,p1,v. subst (lam x s u) p1 x v = lam x s u. +#x #s #u #p1 #v + + +definition subst ≝ λu.λpu:tm u.λx,y. + tm2_ind_plus ? + (* par x0 *) (λx0.match x == x0 with [ true ⇒ v | false ⇒ u ]) + (* app v1 v2 *) (λv1,v2,pv1,pv2,recv1,recv2.app recv1 recv2) + (* lam y#(x::FV v) s w *) (x::FV v) (λy,s,w,py1,py2,pw,recw.lam y s (recw y py1)) + u (tm_to_tm2 … pu). +check subst +definition subst ≝ λu.λpu:tm u.λx,v. + nominal u with + [ xlam y # (x::FV v) s w ^ ? ]. + +(* +notation > "Λ ident x. ident T [ident x] ↦ P" + with precedence 48 for @{'foo (λ${ident x}.λ${ident T}.$P)}. + +notation < "Λ ident x. ident T [ident x] ↦ P" + with precedence 48 for @{'foo (λ${ident x}:$Q.λ${ident T}:$R.$P)}. +*) + +(* +notation +"hvbox('nominal' u 'with' + [ 'xpar' ident x ⇒ f1 + | 'xapp' ident v1 ident v2 ⇒ f2 + | 'xlam' ident x # C s w ⇒ f3 ])" +with precedence 48 +for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1) + (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) + $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. +*) + +(* +notation +"hvbox('nominal' u 'with' + [ 'xpar' ident x ^ f1 + | 'xapp' ident v1 ident v2 ^ f2 ])" +(* | 'xlam' ident x # C s w ^ f3 ]) *) +with precedence 48 +for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1) + (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) + $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. +*) +notation +"hvbox('nominal' u 'with' + [ 'xpar' ident x ^ f1 + | 'xapp' ident v1 ident v2 ^ f2 ])" +with precedence 48 +for @{ tm2_ind_plus ? (λ${ident x}:?.$f1) + (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) + $C (λ${ident x}:?.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. + + +definition subst ≝ λu.λpu:tm u.λx,v. + nominal u with + [ xpar x0 ^ match x == x0 with [ true ⇒ v | false ⇒ u ] + | xapp v1 v2 ^ ? ]. + | xlam y # (x::FV v) s w ^ ? ]. + + + (* par x0 *) (λx0.match x == x0 with [ true ⇒ v | false ⇒ u ]) + (* app v1 v2 *) (λv1,v2,pv1,pv2,recv1,recv2.app recv1 recv2) + (* lam y#(x::FV v) s w *) (x::FV v) (λy,s,w,py1,py2,pw,recw.lam y s (recw y py1)) + u (tm_to_tm2 … pu). + + +*) +definition subst ≝ λu.λpu:tm u.λx,v. + tm2_ind_plus ? + (* par x0 *) (λx0.match x == x0 with [ true ⇒ v | false ⇒ u ]) + (* app v1 v2 *) (λv1,v2,pv1,pv2,recv1,recv2.app recv1 recv2) + (* lam y#(x::FV v) s w *) (x::FV v) (λy,s,w,py1,py2,pw,recw.lam y s (recw y py1)) + u (tm_to_tm2 … pu). + +check subst + + +axiom in_Env : 𝔸 × tp → Env → Prop. +notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}. +interpretation "Env membership" 'lefttriangle x l = (in_Env x l). + + + +inductive judg : list tp → tm → tp → Prop ≝ +| t_var : ∀g,n,t.Nth ? n g = Some ? t → judg g (var n) t +| t_app : ∀g,m,n,t,u.judg g m (arr t u) → judg g n t → judg g (app m n) u +| t_abs : ∀g,t,m,u.judg (t::g) m u → judg g (abs t m) (arr t u). + +definition Env := list (𝔸 × tp). + +axiom vclose_env : Env → list tp. +axiom vclose_tm : Env → tm → tm. +axiom Lam : 𝔸 → tp → tm → tm. +definition Judg ≝ λG,M,T.judg (vclose_env G) (vclose_tm G M) T. +definition dom ≝ λG:Env.map ?? (fst ??) G. + +definition sctx ≝ 𝔸 × tm. +axiom swap_tm : 𝔸 → 𝔸 → tm → tm. +definition sctx_app : sctx → 𝔸 → tm ≝ λM0,Y.let 〈X,M〉 ≝ M0 in swap_tm X Y M. + +axiom in_list : ∀A:Type[0].A → list A → Prop. +interpretation "list membership" 'mem x l = (in_list ? x l). +interpretation "list non-membership" 'notmem x l = (Not (in_list ? x l)). + +axiom in_Env : 𝔸 × tp → Env → Prop. +notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}. +interpretation "Env membership" 'lefttriangle x l = (in_Env x l). + +let rec FV M ≝ match M with + [ par X ⇒ [X] + | app M1 M2 ⇒ FV M1@FV M2 + | abs T M0 ⇒ FV M0 + | _ ⇒ [ ] ]. + +(* axiom Lookup : 𝔸 → Env → option tp. *) + +(* forma alto livello del judgment + t_abs* : ∀G,T,X,M,U. + (∀Y ∉ supp(M).Judg (〈Y,T〉::G) (M[Y]) U) → + Judg G (Lam X T (M[X])) (arr T U) *) + +(* prima dimostrare, poi perfezionare gli assiomi, poi dimostrarli *) + +axiom Judg_ind : ∀P:Env → tm → tp → Prop. + (∀X,G,T.〈X,T〉 ◃ G → P G (par X) T) → + (∀G,M,N,T,U. + Judg G M (arr T U) → Judg G N T → + P G M (arr T U) → P G N T → P G (app M N) U) → + (∀G,T1,T2,X,M1. + (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → Judg (〈Y,T1〉::G) (sctx_app M1 Y) T2) → + (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → P (〈Y,T1〉::G) (sctx_app M1 Y) T2) → + P G (Lam X T1 (sctx_app M1 X)) (arr T1 T2)) → + ∀G,M,T.Judg G M T → P G M T. + +axiom t_par : ∀X,G,T.〈X,T〉 ◃ G → Judg G (par X) T. +axiom t_app2 : ∀G,M,N,T,U.Judg G M (arr T U) → Judg G N T → Judg G (app M N) U. +axiom t_Lam : ∀G,X,M,T,U.Judg (〈X,T〉::G) M U → Judg G (Lam X T M) (arr T U). + +definition subenv ≝ λG1,G2.∀x.x ◃ G1 → x ◃ G2. +interpretation "subenv" 'subseteq G1 G2 = (subenv G1 G2). + +axiom daemon : ∀P:Prop.P. + +theorem weakening : ∀G1,G2,M,T.G1 ⊆ G2 → Judg G1 M T → Judg G2 M T. +#G1 #G2 #M #T #Hsub #HJ lapply Hsub lapply G2 -G2 change with (∀G2.?) +@(Judg_ind … HJ) +[ #X #G #T0 #Hin #G2 #Hsub @t_par @Hsub // +| #G #M0 #N #T0 #U #HM0 #HN #IH1 #IH2 #G2 #Hsub @t_app2 + [| @IH1 // | @IH2 // ] +| #G #T1 #T2 #X #M1 #HM1 #IH #G2 #Hsub @t_Lam @IH + [ (* trivial property of Lam *) @daemon + | (* trivial property of subenv *) @daemon ] +] +qed. + +(* Serve un tipo Tm per i termini localmente chiusi e i suoi principi di induzione e + ricorsione *) \ No newline at end of file diff --git a/matita/matita/broken_lib/binding/names.ma b/matita/matita/broken_lib/binding/names.ma new file mode 100644 index 000000000..bc5f07d43 --- /dev/null +++ b/matita/matita/broken_lib/binding/names.ma @@ -0,0 +1,78 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basics/logic.ma". +include "basics/lists/in.ma". +include "basics/types.ma". + +(*interpretation "list membership" 'mem x l = (in_list ? x l).*) + +record Nset : Type[1] ≝ +{ + (* carrier is specified as a coercion: when an object X of type Nset is + given, but something of type Type is expected, Matita will insert a + hidden coercion: the user sees "X", but really means "carrier X" *) + carrier :> DeqSet; + N_fresh : list carrier → carrier; + p_fresh : ∀l.N_fresh l ∉ l +}. + +definition maxlist ≝ + λl.foldr ?? (λx,acc.max x acc) 0 l. + +definition natfresh ≝ λl.S (maxlist l). + +lemma le_max_1 : ∀x,y.x ≤ max x y. /2/ +qed. + +lemma le_max_2 : ∀x,y.y ≤ max x y. /2/ +qed. + +lemma le_maxlist : ∀l,x.x ∈ l → x ≤ maxlist l. +#l elim l +[#x #Hx @False_ind cases (not_in_list_nil ? x) #H1 /2/ +|#y #tl #IH #x #H1 change with (max ??) in ⊢ (??%); + cases (in_list_cons_case ???? H1);#H2; + [ >H2 @le_max_1 + | whd in ⊢ (??%); lapply (refl ? (leb y (maxlist tl))); + cases (leb y (maxlist tl)) in ⊢ (???% → %);#H3 + [ @IH // + | lapply (IH ? H2) #H4 + lapply (leb_false_to_not_le … H3) #H5 + lapply (not_le_to_lt … H5) #H6 + @(transitive_le … H4) + @(transitive_le … H6) %2 % + ] + ] +] +qed. + +(* prove freshness for nat *) +lemma lt_l_natfresh_l : ∀l,x.x ∈ l → x < natfresh l. +#l #x #H1 @le_S_S /2/ +qed. + +(*naxiom p_Xfresh : ∀l.∀x:Xcarr.x ∈ l → x ≠ ntm (Xfresh l) ∧ x ≠ ntp (Xfresh l).*) +lemma p_natfresh : ∀l.natfresh l ∉ l. +#l % #H1 lapply (lt_l_natfresh_l … H1) #H2 +cases (lt_to_not_eq … H2) #H3 @H3 % +qed. + +include "basics/finset.ma". + +definition X : Nset ≝ mk_Nset DeqNat …. +[ @natfresh +| @p_natfresh +] +qed. \ No newline at end of file diff --git a/matita/matita/broken_lib/finite_lambda/confluence.ma b/matita/matita/broken_lib/finite_lambda/confluence.ma new file mode 100644 index 000000000..32b7e3ff0 --- /dev/null +++ b/matita/matita/broken_lib/finite_lambda/confluence.ma @@ -0,0 +1,226 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "finite_lambda/reduction.ma". + + +axiom canonical_to_T: ∀O,D.∀M:T O D.∀ty.(* type_of M ty → *) + ∃a:FinSet_of_FType O D ty. star ? (red O D) M (to_T O D ty a). + +axiom normal_to_T: ∀O,D,M,ty,a. red O D (to_T O D ty a) M → False. + +axiom red_closed: ∀O,D,M,M1. + is_closed O D 0 M → red O D M M1 → is_closed O D 0 M1. + +lemma critical: ∀O,D,ty,M,N. + ∃M3:T O D + .star (T O D) (red O D) (subst O D M 0 N) M3 + ∧star (T O D) (red O D) + (App O D + (Vec O D ty + (map (FinSet_of_FType O D ty) (T O D) + (λa0:FinSet_of_FType O D ty.subst O D M 0 (to_T O D ty a0)) + (enum (FinSet_of_FType O D ty)))) N) M3. +#O #D #ty #M #N +lapply (canonical_to_T O D N ty) * #a #Ha +%{(subst O D M 0 (to_T O D ty a))} (* CR-term *) +%[@red_star_subst @Ha + |@trans_star [|@(star_red_appr … Ha)] @R_to_star @riota + lapply (enum_complete (FinSet_of_FType O D ty) a) + elim (enum (FinSet_of_FType O D ty)) + [normalize #H1 destruct (H1) + |#hd #tl #Hind #H cases (orb_true_l … H) -H #Hcase + [normalize >Hcase >(\P Hcase) // + |normalize cases (true_or_false (a==hd)) #Hcase1 + [normalize >Hcase1 >(\P Hcase1) // |>Hcase1 @Hind @Hcase] + ] + ] + ] +qed. + +lemma critical2: ∀O,D,ty,a,M,M1,M2,v. + red O D (Vec O D ty v) M → + red O D (App O D (Vec O D ty v) (to_T O D ty a)) M1 → + assoc (FinSet_of_FType O D ty) (T O D) a (enum (FinSet_of_FType O D ty)) v + =Some (T O D) M2 → + ∃M3:T O D + .star (T O D) (red O D) M2 M3 + ∧star (T O D) (red O D) (App O D M (to_T O D ty a)) M3. +#O #D #ty #a #M #M1 #M2 #v #redM #redM1 #Ha lapply (red_vec … redM) -redM +* #N * #N1 * #v1 * #v2 * * #Hred1 #Hv #HM0 >HM0 -HM0 >Hv in Ha; #Ha +cases (same_assoc … a (enum (FinSet_of_FType O D ty)) v1 v2 N N1) + [* >Ha -Ha #H1 destruct (H1) #Ha + %{N1} (* CR-term *) % [@R_to_star //|@R_to_star @(riota … Ha)] + |#Ha1 %{M2} (* CR-term *) % [// | @R_to_star @riota length_map >length_map //] #n #M0 + cases (true_or_false (leb (|enum (FinSet_of_FType O D ty)|) n)) #Hcase + [>nth_to_default [2:>length_map @(leb_true_to_le … Hcase)] + >nth_to_default [2:>length_map @(leb_true_to_le … Hcase)] // + |cut (n < |enum (FinSet_of_FType O D ty)|) + [@not_le_to_lt @leb_false_to_not_le @Hcase] #Hlt + cut (∃a:FinSet_of_FType O D ty.True) + [lapply Hlt lapply (enum_complete (FinSet_of_FType O D ty)) + cases (enum (FinSet_of_FType O D ty)) + [#_ normalize #H @False_ind @(absurd … H) @lt_to_not_le // + |#a #l #_ #_ %{a} // + ] + ] * #a #_ + >(nth_map ?????? a Hlt) >(nth_map ?????? a Hlt) #_ + @red_star_subst2 // + ] + ] +qed. + +(* we need to proceed by structural induction on the term and then +by inversion on the two redexes. The problem are the moves in a +same subterm, since we need an induction hypothesis, there *) + +lemma local_confluence: ∀O,D,M,M1,M2. red O D M M1 → red O D M M2 → +∃M3. star ? (red O D) M1 M3 ∧ star ? (red O D) M2 M3. +#O #D #M @(T_elim … M) + [#o #a #M1 #M2 #H elim(red_val ????? H) + |#n #M1 #M2 #H elim(red_rel ???? H) + |(* app : this is the interesting case *) + #P #Q #HindP #HindQ + #M1 #M2 #H1 inversion H1 -H1 + [(* right redex is beta *) + #ty #Q #N #Hc #HM >HM -HM #HM1 >HM1 - HM1 #Hl inversion Hl + [#ty1 #Q1 #N1 #Hc1 #H1 destruct (H1) #H_ + %{(subst O D Q1 0 N1)} (* CR-term *) /2/ + |#ty #v #a #M0 #_ #H1 destruct (H1) (* vacuous *) + |#M0 #M10 #N0 #redM0 #_ #H1 destruct (H1) #_ cases (red_lambda … redM0) + [* #Q1 * #redQ #HM10 >HM10 + %{(subst O D Q1 0 N0)} (* CR-term *) % + [@red_star_subst2 //|@R_to_star @rbeta @Hc] + |#HM1 >HM1 @critical + ] + |#M0 #N0 #N1 #redN0N1 #_ #H1 destruct (H1) #HM2 + %{(subst O D Q 0 N1)} (* CR-term *) + %[@red_star_subst @R_to_star //|@R_to_star @rbeta @(red_closed … Hc) //] + |#ty1 #N0 #N1 #_ #_ #H1 destruct (H1) (* vacuous *) + |#ty1 #M0 #H1 destruct (H1) (* vacuous *) + |#ty1 #N0 #N1 #v #v1 #_ #_ #H1 destruct (H1) (* vacuous *) + ] + |(* right redex is iota *)#ty #v #a #M3 #Ha #_ #_ #Hl inversion Hl + [#P1 #M1 #N1 #_ #H1 destruct (H1) (* vacuous *) + |#ty1 #v1 #a1 #M4 #Ha1 #H1 destruct (H1) -H1 #HM4 >(inj_to_T … e0) in Ha; + >Ha1 #H1 destruct (H1) %{M3} (* CR-term *) /2/ + |#M0 #M10 #N0 #redM0 #_ #H1 destruct (H1) #HM2 @(critical2 … redM0 Hl Ha) + |#M0 #N0 #N1 #redN0N1 #_ #H1 destruct (H1) elim (normal_to_T … redN0N1) + |#ty1 #N0 #N1 #_ #_ #H1 destruct (H1) (* vacuous *) + |#ty1 #M0 #H1 destruct (H1) (* vacuous *) + |#ty1 #N0 #N1 #v #v1 #_ #_ #H1 destruct (H1) (* vacuous *) + ] + |(* right redex is appl *)#M3 #M4 #N #redM3M4 #_ #H1 destruct (H1) #_ + #Hl inversion Hl + [#ty1 #M1 #N1 #Hc #H1 destruct (H1) #HM2 lapply (red_lambda … redM3M4) * + [* #M3 * #H1 #H2 >H2 %{(subst O D M3 0 N1)} % + [@R_to_star @rbeta @Hc|@red_star_subst2 // ] + |#H >H -H lapply (critical O D ty1 M1 N1) * #M3 * #H1 #H2 + %{M3} /2/ + ] + |#ty1 #v1 #a1 #M4 #Ha1 #H1 #H2 destruct + lapply (critical2 … redM3M4 Hl Ha1) * #M3 * #H1 #H2 %{M3} /2/ + |#M0 #M10 #N0 #redM0 #_ #H1 destruct (H1) #HM2 + lapply (HindP … redM0 redM3M4) * #M3 * #H1 #H2 + %{(App O D M3 N0)} (* CR-term *) % [@star_red_appl //|@star_red_appl //] + |#M0 #N0 #N1 #redN0N1 #_ #H1 destruct (H1) #_ + %{(App O D M4 N1)} % @R_to_star [@rappr //|@rappl //] + |#ty1 #N0 #N1 #_ #_ #H1 destruct (H1) (* vacuous *) + |#ty1 #M0 #H1 destruct (H1) (* vacuous *) + |#ty1 #N0 #N1 #v #v1 #_ #_ #H1 destruct (H1) (* vacuous *) + ] + |(* right redex is appr *)#M3 #N #N1 #redN #_ #H1 destruct (H1) #_ + #Hl inversion Hl + [#ty1 #M0 #N0 #Hc #H1 destruct (H1) #HM2 + %{(subst O D M0 0 N1)} (* CR-term *) % + [@R_to_star @rbeta @(red_closed … Hc) //|@red_star_subst @R_to_star // ] + |#ty1 #v1 #a1 #M4 #Ha1 #H1 #H2 destruct (H1) elim (normal_to_T … redN) + |#M0 #M10 #N0 #redM0 #_ #H1 destruct (H1) #HM2 + %{(App O D M10 N1)} (* CR-term *) % @R_to_star [@rappl //|@rappr //] + |#M0 #N0 #N10 #redN0 #_ #H1 destruct (H1) #_ + lapply (HindQ … redN0 redN) * #M3 * #H1 #H2 + %{(App O D M0 M3)} (* CR-term *) % [@star_red_appr //|@star_red_appr //] + |#ty1 #N0 #N1 #_ #_ #H1 destruct (H1) (* vacuous *) + |#ty1 #M0 #H1 destruct (H1) (* vacuous *) + |#ty1 #N0 #N1 #v #v1 #_ #_ #H1 destruct (H1) (* vacuous *) + ] + |(* right redex is rlam *) #ty #N0 #N1 #_ #_ #H1 destruct (H1) (* vacuous *) + |(* right redex is rmem *) #ty #M0 #H1 destruct (H1) (* vacuous *) + |(* right redex is vec *) #ty #N #N1 #v #v1 #_ #_ + #H1 destruct (H1) (* vacuous *) + ] + |#ty #M1 #Hind #M2 #M3 #H1 #H2 (* this case is not trivial any more *) + lapply (red_lambda … H1) * + [* #M4 * #H3 #H4 >H4 lapply (red_lambda … H2) * + [* #M5 * #H5 #H6 >H6 lapply(Hind … H3 H5) * #M6 * #H7 #H8 + %{(Lambda O D ty M6)} (* CR-term *) % @star_red_lambda // + |#H5 >H5 @critical3 // + ] + |#HM2 >HM2 lapply (red_lambda … H2) * + [* #M4 * #Hred #HM3 >HM3 lapply (critical3 … ty ?? Hred) * #M5 + * #H3 #H4 %{M5} (* CR-term *) % // + |#HM3 >HM3 %{M3} (* CR-term *) % // + ] + ] + |#ty #v1 #Hind #M1 #M2 #H1 #H2 + lapply (red_vec … H1) * #N11 * #N12 * #v11 * #v12 * * #redN11 #Hv1 #HM1 + lapply (red_vec … H2) * #N21* #N22 * #v21 * #v22 * * #redN21 #Hv2 #HM2 + >Hv1 in Hv2; #Hvv lapply (compare_append … Hvv) -Hvv * + (* we must proceed by cases on the list *) * normalize + [(* N11 = N21 *) * + [>append_nil * #Hl1 #Hl2 destruct lapply(Hind N11 … redN11 redN21) + [@mem_append_l2 %1 //] + * #M3 * #HM31 #HM32 + %{(Vec O D ty (v21@M3::v12))} (* CR-term *) + % [@star_red_vec //|@star_red_vec //] + |>append_nil * #Hl1 #Hl2 destruct lapply(Hind N21 … redN21 redN11) + [@mem_append_l2 %1 //] + * #M3 * #HM31 #HM32 + %{(Vec O D ty (v11@M3::v22))} (* CR-term *) + % [@star_red_vec //|@star_red_vec //] + ] + |(* N11 ≠ N21 *) -Hind #P #l * + [* #Hv11 #Hv22 destruct + %{((Vec O D ty ((v21@N22::l)@N12::v12)))} (* CR-term *) % @R_to_star + [>associative_append >associative_append normalize @rvec // + |>append_cons append_cons associative_append >associative_append normalize @rvec // + ] + ] + ] + ] +qed. + + + + diff --git a/matita/matita/broken_lib/finite_lambda/reduction.ma b/matita/matita/broken_lib/finite_lambda/reduction.ma new file mode 100644 index 000000000..98c56e10a --- /dev/null +++ b/matita/matita/broken_lib/finite_lambda/reduction.ma @@ -0,0 +1,308 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "finite_lambda/terms_and_types.ma". + +(* some auxiliary lemmas *) + +lemma nth_to_default: ∀A,l,n,d. + |l| ≤ n → nth n A l d = d. +#A #l elim l [//] #a #tl #Hind #n cases n + [#d normalize #H @False_ind @(absurd … H) @lt_to_not_le // + |#m #d normalize #H @Hind @le_S_S_to_le @H + ] +qed. + +lemma mem_nth: ∀A,l,n,d. + n < |l| → mem ? (nth n A l d) l. +#A #l elim l + [#n #d normalize #H @False_ind @(absurd … H) @lt_to_not_le // + |#a #tl #Hind * normalize + [#_ #_ %1 //| #m #d #HSS %2 @Hind @le_S_S_to_le @HSS] + ] +qed. + +lemma nth_map: ∀A,B,l,f,n,d1,d2. + n < |l| → nth n B (map … f l) d1 = f (nth n A l d2). +#n #B #l #f elim l + [#m #d1 #d2 normalize #H @False_ind @(absurd … H) @lt_to_not_le // + |#a #tl #Hind #m #d1 #d2 cases m normalize // + #m1 #H @Hind @le_S_S_to_le @H + ] +qed. + + + +(* end of auxiliary lemmas *) + +let rec to_T O D ty on ty: FinSet_of_FType O D ty → T O D ≝ + match ty return (λty.FinSet_of_FType O D ty → T O D) with + [atom o ⇒ λa.Val O D o a + |arrow ty1 ty2 ⇒ λa:FinFun ??.Vec O D ty1 + (map ((FinSet_of_FType O D ty1)×(FinSet_of_FType O D ty2)) + (T O D) (λp.to_T O D ty2 (snd … p)) (pi1 … a)) + ] +. + +lemma is_closed_to_T: ∀O,D,ty,a. is_closed O D 0 (to_T O D ty a). +#O #D #ty elim ty // +#ty1 #ty2 #Hind1 #Hind2 #a normalize @cvec #m #Hmem +lapply (mem_map ????? Hmem) * #a1 * #H1 #H2

H //] -Hl1 -Hl2 + lapply e0 -e0 lapply l2 -l2 elim l1 + [#l2 cases l2 normalize [// |#a1 #tl1 #H destruct] + |#a1 #tl1 #Hind #l2 cases l2 + [normalize #H destruct + |#a2 #tl2 normalize #H @eq_f2 + [@Hind2 *) + +let rec assoc (A:FinSet) (B:Type[0]) (a:A) l1 l2 on l1 : option B ≝ + match l1 with + [ nil ⇒ None ? + | cons hd1 tl1 ⇒ match l2 with + [ nil ⇒ None ? + | cons hd2 tl2 ⇒ if a==hd1 then Some ? hd2 else assoc A B a tl1 tl2 + ] + ]. + +lemma same_assoc: ∀A,B,a,l1,v1,v2,N,N1. + assoc A B a l1 (v1@N::v2) = Some ? N ∧ assoc A B a l1 (v1@N1::v2) = Some ? N1 + ∨ assoc A B a l1 (v1@N::v2) = assoc A B a l1 (v1@N1::v2). +#A #B #a #l1 #v1 #v2 #N #N1 lapply v1 -v1 elim l1 + [#v1 %2 // |#hd #tl #Hind * normalize cases (a==hd) normalize /3/] +qed. + +lemma assoc_to_mem: ∀A,B,a,l1,l2,b. + assoc A B a l1 l2 = Some ? b → mem ? b l2. +#A #B #a #l1 elim l1 + [#l2 #b normalize #H destruct + |#hd1 #tl1 #Hind * + [#b normalize #H destruct + |#hd2 #tl2 #b normalize cases (a==hd1) normalize + [#H %1 destruct //|#H %2 @Hind @H] + ] + ] +qed. + +lemma assoc_to_mem2: ∀A,B,a,l1,l2,b. + assoc A B a l1 l2 = Some ? b → ∃l21,l22.l2=l21@b::l22. +#A #B #a #l1 elim l1 + [#l2 #b normalize #H destruct + |#hd1 #tl1 #Hind * + [#b normalize #H destruct + |#hd2 #tl2 #b normalize cases (a==hd1) normalize + [#H %{[]} %{tl2} destruct // + |#H lapply (Hind … H) * #la * #lb #H1 + %{(hd2::la)} %{lb} >H1 //] + ] + ] +qed. + +lemma assoc_map: ∀A,B,C,a,l1,l2,f,b. + assoc A B a l1 l2 = Some ? b → assoc A C a l1 (map ?? f l2) = Some ? (f b). +#A #B #C #a #l1 elim l1 + [#l2 #f #b normalize #H destruct + |#hd1 #tl1 #Hind * + [#f #b normalize #H destruct + |#hd2 #tl2 #f #b normalize cases (a==hd1) normalize + [#H destruct // |#H @(Hind … H)] + ] + ] +qed. + +(*************************** One step reduction *******************************) + +inductive red (O:Type[0]) (D:O→FinSet) : T O D →T O D → Prop ≝ + | (* we only allow beta on closed arguments *) + rbeta: ∀P,M,N. is_closed O D 0 N → + red O D (App O D (Lambda O D P M) N) (subst O D M 0 N) + | riota: ∀ty,v,a,M. + assoc ?? a (enum (FinSet_of_FType O D ty)) v = Some ? M → + red O D (App O D (Vec O D ty v) (to_T O D ty a)) M + | rappl: ∀M,M1,N. red O D M M1 → red O D (App O D M N) (App O D M1 N) + | rappr: ∀M,N,N1. red O D N N1 → red O D (App O D M N) (App O D M N1) + | rlam: ∀ty,N,N1. red O D N N1 → red O D (Lambda O D ty N) (Lambda O D ty N1) + | rmem: ∀ty,M. red O D (Lambda O D ty M) + (Vec O D ty (map ?? (λa. subst O D M 0 (to_T O D ty a)) + (enum (FinSet_of_FType O D ty)))) + | rvec: ∀ty,N,N1,v,v1. red O D N N1 → + red O D (Vec O D ty (v@N::v1)) (Vec O D ty (v@N1::v1)). + +(*********************************** inversion ********************************) +lemma red_vec: ∀O,D,ty,v,M. + red O D (Vec O D ty v) M → ∃N,N1,v1,v2. + red O D N N1 ∧ v = v1@N::v2 ∧ M = Vec O D ty (v1@N1::v2). +#O #D #ty #v #M #Hred inversion Hred + [#ty1 #M0 #N #Hc #H destruct + |#ty1 #v1 #a #M0 #_ #H destruct + |#M0 #M1 #N #_ #_ #H destruct + |#M0 #M1 #N #_ #_ #H destruct + |#ty1 #M #M1 #_ #_ #H destruct + |#ty1 #M0 #H destruct + |#ty1 #N #N1 #v1 #v2 #Hred1 #_ #H destruct #_ %{N} %{N1} %{v1} %{v2} /3/ + ] +qed. + +lemma red_lambda: ∀O,D,ty,M,N. + red O D (Lambda O D ty M) N → + (∃M1. red O D M M1 ∧ N = (Lambda O D ty M1)) ∨ + N = Vec O D ty (map ?? (λa. subst O D M 0 (to_T O D ty a)) + (enum (FinSet_of_FType O D ty))). +#O #D #ty #M #N #Hred inversion Hred + [#ty1 #M0 #N #Hc #H destruct + |#ty1 #v1 #a #M0 #_ #H destruct + |#M0 #M1 #N #_ #_ #H destruct + |#M0 #M1 #N #_ #_ #H destruct + |#ty1 #P #P1 #redP #_ #H #H1 destruct %1 %{P1} % // + |#ty1 #M0 #H destruct #_ %2 // + |#ty1 #N #N1 #v1 #v2 #Hred1 #_ #H destruct + ] +qed. + +lemma red_val: ∀O,D,ty,a,N. + red O D (Val O D ty a) N → False. +#O #D #ty #M #N #Hred inversion Hred + [#ty1 #M0 #N #Hc #H destruct + |#ty1 #v1 #a #M0 #_ #H destruct + |#M0 #M1 #N #_ #_ #H destruct + |#M0 #M1 #N #_ #_ #H destruct + |#ty1 #N1 #N2 #_ #_ #H destruct + |#ty1 #M0 #H destruct #_ + |#ty1 #N #N1 #v1 #v2 #Hred1 #_ #H destruct + ] +qed. + +lemma red_rel: ∀O,D,n,N. + red O D (Rel O D n) N → False. +#O #D #n #N #Hred inversion Hred + [#ty1 #M0 #N #Hc #H destruct + |#ty1 #v1 #a #M0 #_ #H destruct + |#M0 #M1 #N #_ #_ #H destruct + |#M0 #M1 #N #_ #_ #H destruct + |#ty1 #N1 #N2 #_ #_ #H destruct + |#ty1 #M0 #H destruct #_ + |#ty1 #N #N1 #v1 #v2 #Hred1 #_ #H destruct + ] +qed. + +(*************************** multi step reduction *****************************) +lemma star_red_appl: ∀O,D,M,M1,N. star ? (red O D) M M1 → + star ? (red O D) (App O D M N) (App O D M1 N). +#O #D #M #N #N1 #H elim H // +#P #Q #Hind #HPQ #Happ %1[|@Happ] @rappl @HPQ +qed. + +lemma star_red_appr: ∀O,D,M,N,N1. star ? (red O D) N N1 → + star ? (red O D) (App O D M N) (App O D M N1). +#O #D #M #N #N1 #H elim H // +#P #Q #Hind #HPQ #Happ %1[|@Happ] @rappr @HPQ +qed. + +lemma star_red_vec: ∀O,D,ty,N,N1,v1,v2. star ? (red O D) N N1 → + star ? (red O D) (Vec O D ty (v1@N::v2)) (Vec O D ty (v1@N1::v2)). +#O #D #ty #N #N1 #v1 #v2 #H elim H // +#P #Q #Hind #HPQ #Hvec %1[|@Hvec] @rvec @HPQ +qed. + +lemma star_red_vec1: ∀O,D,ty,v1,v2,v. |v1| = |v2| → + (∀n,M. n < |v1| → star ? (red O D) (nth n ? v1 M) (nth n ? v2 M)) → + star ? (red O D) (Vec O D ty (v@v1)) (Vec O D ty (v@v2)). +#O #D #ty #v1 elim v1 + [#v2 #v normalize #Hv2 >(lenght_to_nil … (sym_eq … Hv2)) normalize // + |#N1 #tl1 #Hind * [normalize #v #H destruct] #N2 #tl2 #v normalize #HS + #H @(trans_star … (Vec O D ty (v@N2::tl1))) + [@star_red_vec @(H 0 N1) @le_S_S // + |>append_cons >(append_cons ??? tl2) @(Hind… (injective_S … HS)) + #n #M #H1 @(H (S n)) @le_S_S @H1 + ] + ] +qed. + +lemma star_red_vec2: ∀O,D,ty,v1,v2. |v1| = |v2| → + (∀n,M. n < |v1| → star ? (red O D) (nth n ? v1 M) (nth n ? v2 M)) → + star ? (red O D) (Vec O D ty v1) (Vec O D ty v2). +#O #D #ty #v1 #v2 @(star_red_vec1 … [ ]) +qed. + +lemma star_red_lambda: ∀O,D,ty,N,N1. star ? (red O D) N N1 → + star ? (red O D) (Lambda O D ty N) (Lambda O D ty N1). +#O #D #ty #N #N1 #H elim H // +#P #Q #Hind #HPQ #Hlam %1[|@Hlam] @rlam @HPQ +qed. + +(************************ reduction and substitution **************************) + +lemma red_star_subst : ∀O,D,M,N,N1,i. + star ? (red O D) N N1 → star ? (red O D) (subst O D M i N) (subst O D M i N1). +#O #D #M #N #N1 #i #Hred lapply i -i @(T_elim … M) normalize + [#o #a #i // + |#i #n cases (leb n i) normalize // cases (eqb n i) normalize // + |#P #Q #HindP #HindQ #n normalize + @(trans_star … (App O D (subst O D P n N1) (subst O D Q n N))) + [@star_red_appl @HindP |@star_red_appr @HindQ] + |#ty #P #HindP #i @star_red_lambda @HindP + |#ty #v #Hindv #i @star_red_vec2 [>length_map >length_map //] + #j #Q inversion v [#_ normalize //] #a #tl #_ #Hv + cases (true_or_false (leb (S j) (|a::tl|))) #Hcase + [lapply (leb_true_to_le … Hcase) -Hcase #Hcase + >(nth_map ?????? a Hcase) >(nth_map ?????? a Hcase) #_ @Hindv >Hv @mem_nth // + |>nth_to_default + [2:>length_map @le_S_S_to_le @not_le_to_lt @leb_false_to_not_le //] + >nth_to_default + [2:>length_map @le_S_S_to_le @not_le_to_lt @leb_false_to_not_le //] // + ] + ] +qed. + +lemma red_star_subst2 : ∀O,D,M,M1,N,i. is_closed O D 0 N → + red O D M M1 → star ? (red O D) (subst O D M i N) (subst O D M1 i N). +#O #D #M #M1 #N #i #HNc #Hred lapply i -i elim Hred + [#ty #P #Q #HQc #i normalize @starl_to_star @sstepl + [|@rbeta >(subst_closed … HQc) //] >(subst_closed … HQc) // + lapply (subst_lemma ?? P ?? i 0 (is_closed_mono … HQc) HNc) // + (subst_closed … (le_O_n …)) // + @R_to_star @riota @assoc_map @HP + |#P #P1 #Q #Hred #Hind #i normalize @star_red_appl @Hind + |#P #P1 #Q #Hred #Hind #i normalize @star_red_appr @Hind + |#ty #P #P1 #Hred #Hind #i normalize @star_red_lambda @Hind + |#ty #P #i normalize @starl_to_star @sstepl [|@rmem] + @star_to_starl @star_red_vec2 [>length_map >length_map >length_map //] + #n #Q >length_map #H + cut (∃a:(FinSet_of_FType O D ty).True) + [lapply H -H lapply (enum_complete (FinSet_of_FType O D ty)) + cases (enum (FinSet_of_FType O D ty)) + [#x normalize #H @False_ind @(absurd … H) @lt_to_not_le // + |#x #l #_ #_ %{x} // + ] + ] * #a #_ + >(nth_map ?????? a H) >(nth_map ?????? Q) [2:>length_map @H] + >(nth_map ?????? a H) + lapply (subst_lemma O D P (to_T O D ty + (nth n (FinSet_of_FType O D ty) (enum (FinSet_of_FType O D ty)) a)) + N i 0 (is_closed_mono … (is_closed_to_T …)) HNc) // H1 + Hx @Hind2 normalize // + |@Hind1 #N #H @Hind2 @(lt_to_le_to_lt … H) normalize // + ] + ] + ] +qed. + +(* since we only consider beta reduction with closed arguments we could avoid +lifting. We define it for the sake of generality *) + +(* arguments: k is the nesting depth (starts from 0), p is the lift +let rec lift O D t k p on t ≝ + match t with + [ Val o a ⇒ Val O D o a + | Rel n ⇒ if (leb k n) then Rel O D (n+p) else Rel O D n + | App m n ⇒ App O D (lift O D m k p) (lift O D n k p) + | Lambda Ty n ⇒ Lambda O D Ty (lift O D n (S k) p) + | Vec Ty v ⇒ Vec O D Ty (map ?? (λx. lift O D x k p) v) + ]. + +notation "↑ ^ n ( M )" non associative with precedence 40 for @{'Lift 0 $n $M}. +notation "↑ _ k ^ n ( M )" non associative with precedence 40 for @{'Lift $n $k $M}. + +interpretation "Lift" 'Lift n k M = (lift ?? M k n). + +let rec subst O D t k s on t ≝ + match t with + [ Val o a ⇒ Val O D o a + | Rel n ⇒ if (leb k n) then + (if (eqb k n) then lift O D s 0 n else Rel O D (n-1)) + else(Rel O D n) + | App m n ⇒ App O D (subst O D m k s) (subst O D n k s) + | Lambda T n ⇒ Lambda O D T (subst O D n (S k) s) + | Vec T v ⇒ Vec O D T (map ?? (λx. subst O D x k s) v) + ]. +*) + +(* simplified version of subst, assuming the argument s is closed *) + +let rec subst O D t k s on t ≝ + match t with + [ Val o a ⇒ Val O D o a + | Rel n ⇒ if (leb k n) then + (if (eqb k n) then (* lift O D s 0 n*) s else Rel O D (n-1)) + else(Rel O D n) + | App m n ⇒ App O D (subst O D m k s) (subst O D n k s) + | Lambda T n ⇒ Lambda O D T (subst O D n (S k) s) + | Vec T v ⇒ Vec O D T (map ?? (λx. subst O D x k s) v) + ]. +(* notation "hvbox(M break [ k ≝ N ])" + non associative with precedence 90 + for @{'Subst1 $M $k $N}. *) + +interpretation "Subst" 'Subst1 M k N = (subst M k N). + +(* +lemma subst_rel1: ∀O,D,A.∀k,i. i < k → + (Rel O D i) [k ≝ A] = Rel O D i. +#A #k #i normalize #ltik >(lt_to_leb_false … ltik) // +qed. + +lemma subst_rel2: ∀O,D, A.∀k. + (Rel k) [k ≝ A] = lift A 0 k. +#A #k normalize >(le_to_leb_true k k) // >(eq_to_eqb_true … (refl …)) // +qed. + +lemma subst_rel3: ∀A.∀k,i. k < i → + (Rel i) [k ≝ A] = Rel (i-1). +#A #k #i normalize #ltik >(le_to_leb_true k i) /2/ +>(not_eq_to_eqb_false k i) // @lt_to_not_eq // +qed. *) + + +(* closed terms ???? +let rec closed_k O D (t: T O D) k on t ≝ + match t with + [ Val o a ⇒ True + | Rel n ⇒ n < k + | App m n ⇒ (closed_k O D m k) ∧ (closed_k O D n k) + | Lambda T n ⇒ closed_k O D n (k+1) + | Vec T v ⇒ closed_list O D v k + ] + +and closed_list O D (l: list (T O D)) k on l ≝ + match l with + [ nil ⇒ True + | cons hd tl ⇒ closed_k O D hd k ∧ closed_list O D tl k + ] +. *) + +inductive is_closed (O:Type[0]) (D:O→FinSet): nat → T O D → Prop ≝ +| cval : ∀k,o,a.is_closed O D k (Val O D o a) +| crel : ∀k,n. n < k → is_closed O D k (Rel O D n) +| capp : ∀k,m,n. is_closed O D k m → is_closed O D k n → + is_closed O D k (App O D m n) +| clam : ∀T,k,m. is_closed O D (S k) m → is_closed O D k (Lambda O D T m) +| cvec: ∀T,k,v. (∀m. mem ? m v → is_closed O D k m) → + is_closed O D k (Vec O D T v). + +lemma is_closed_rel: ∀O,D,n,k. + is_closed O D k (Rel O D n) → n < k. +#O #D #n #k #H inversion H + [#k0 #o #a #eqk #H destruct + |#k0 #n0 #ltn0 #eqk #H destruct // + |#k0 #M #N #_ #_ #_ #_ #_ #H destruct + |#T #k0 #M #_ #_ #_ #H destruct + |#T #k0 #v #_ #_ #_ #H destruct + ] +qed. + +lemma is_closed_app: ∀O,D,k,M, N. + is_closed O D k (App O D M N) → is_closed O D k M ∧ is_closed O D k N. +#O #D #k #M #N #H inversion H + [#k0 #o #a #eqk #H destruct + |#k0 #n0 #ltn0 #eqk #H destruct + |#k0 #M1 #N1 #HM #HN #_ #_ #_ #H1 destruct % // + |#T #k0 #M #_ #_ #_ #H destruct + |#T #k0 #v #_ #_ #_ #H destruct + ] +qed. + +lemma is_closed_lam: ∀O,D,k,ty,M. + is_closed O D k (Lambda O D ty M) → is_closed O D (S k) M. +#O #D #k #ty #M #H inversion H + [#k0 #o #a #eqk #H destruct + |#k0 #n0 #ltn0 #eqk #H destruct + |#k0 #M1 #N1 #HM #HN #_ #_ #_ #H1 destruct + |#T #k0 #M1 #HM1 #_ #_ #H1 destruct // + |#T #k0 #v #_ #_ #_ #H destruct + ] +qed. + +lemma is_closed_vec: ∀O,D,k,ty,v. + is_closed O D k (Vec O D ty v) → ∀m. mem ? m v → is_closed O D k m. +#O #D #k #ty #M #H inversion H + [#k0 #o #a #eqk #H destruct + |#k0 #n0 #ltn0 #eqk #H destruct + |#k0 #M1 #N1 #HM #HN #_ #_ #_ #H1 destruct + |#T #k0 #M1 #HM1 #_ #_ #H1 destruct + |#T #k0 #v #Hv #_ #_ #H1 destruct @Hv + ] +qed. + +lemma is_closed_S: ∀O,D,M,m. + is_closed O D m M → is_closed O D (S m) M. +#O #D #M #m #H elim H // + [#k #n0 #Hlt @crel @le_S // + |#k #P #Q #HP #HC #H1 #H2 @capp // + |#ty #k #P #HP #H1 @clam // + |#ty #k #v #Hind #Hv @cvec @Hv + ] +qed. + +lemma is_closed_mono: ∀O,D,M,m,n. m ≤ n → + is_closed O D m M → is_closed O D n M. +#O #D #M #m #n #lemn elim lemn // #i #j #H #H1 @is_closed_S @H @H1 +qed. + + +(*** properties of lift and subst ***) + +(* +lemma lift_0: ∀O,D.∀t:T O D.∀k. lift O D t k 0 = t. +#O #D #t @(T_elim … t) normalize // + [#n #k cases (leb k n) normalize // + |#o #v #Hind #k @eq_f lapply Hind -Hind elim v // + #hd #tl #Hind #Hind1 normalize @eq_f2 + [@Hind1 %1 //|@Hind #x #Hx @Hind1 %2 //] + ] +qed. + +lemma lift_closed: ∀O,D.∀t:T O D.∀k,p. + is_closed O D k t → lift O D t k p = t. +#O #D #t @(T_elim … t) normalize // + [#n #k #p #H >(not_le_to_leb_false … (lt_to_not_le … (is_closed_rel … H))) // + |#M #N #HindM #HindN #k #p #H lapply (is_closed_app … H) * #HcM #HcN + >(HindM … HcM) >(HindN … HcN) // + |#ty #M #HindM #k #p #H lapply (is_closed_lam … H) -H #H >(HindM … H) // + |#ty #v #HindM #k #p #H lapply (is_closed_vec … H) -H #H @eq_f + cut (∀m. mem ? m v → lift O D m k p = m) + [#m #Hmem @HindM [@Hmem | @H @Hmem]] -HindM + elim v // #a #tl #Hind #H1 normalize @eq_f2 + [@H1 %1 //|@Hind #m #Hmem @H1 %2 @Hmem] + ] +qed. + +*) + +lemma subst_closed: ∀O,D,M,N,k,i. k ≤ i → + is_closed O D k M → subst O D M i N = M. +#O #D #M @(T_elim … M) + [#o #a normalize // + |#n #N #k #j #Hlt #Hc lapply (is_closed_rel … Hc) #Hnk normalize + >not_le_to_leb_false [2:@lt_to_not_le @(lt_to_le_to_lt … Hnk Hlt)] // + |#P #Q #HindP #HindQ #N #k #i #ltki #Hc lapply (is_closed_app … Hc) * + #HcP #HcQ normalize >(HindP … ltki HcP) >(HindQ … ltki HcQ) // + |#ty #P #HindP #N #k #i #ltki #Hc lapply (is_closed_lam … Hc) + #HcP normalize >(HindP … HcP) // @le_S_S @ltki + |#ty #v #Hindv #N #k #i #ltki #Hc lapply (is_closed_vec … Hc) + #Hcv normalize @eq_f + cut (∀m:T O D.mem (T O D) m v→ subst O D m i N=m) + [#m #Hmem @(Hindv … Hmem N … ltki) @Hcv @Hmem] + elim v // #a #tl #Hind #H normalize @eq_f2 + [@H %1 //| @Hind #Hmem #Htl @H %2 @Htl] + ] +qed. + +lemma subst_lemma: ∀O,D,A,B,C,k,i. is_closed O D k B → is_closed O D i C → + subst O D (subst O D A i B) (k+i) C = + subst O D (subst O D A (k+S i) C) i B. +#O #D #A #B #C #k @(T_elim … A) normalize + [// + |#n #i #HBc #HCc @(leb_elim i n) #Hle + [@(eqb_elim i n) #eqni + [(lt_to_leb_false (k+(S i)) i) // normalize + >(subst_closed … HBc) // >le_to_leb_true // >eq_to_eqb_true // + |(cut (i < n)) + [cases (le_to_or_lt_eq … Hle) // #eqin @False_ind /2/] #ltin + (cut (0 < n)) [@(le_to_lt_to_lt … ltin) //] #posn + normalize @(leb_elim (k+i) (n-1)) #nk + [@(eqb_elim (k+i) (n-1)) #H normalize + [cut (k+(S i) = n); [/2 by S_pred/] #H1 + >(le_to_leb_true (k+(S i)) n) /2/ + >(eq_to_eqb_true … H1) normalize >(subst_closed … HCc) // + |(cut (k+i < n-1)) [@not_eq_to_le_to_lt; //] #Hlt + >(le_to_leb_true (k+(S i)) n) normalize + [>(not_eq_to_eqb_false (k+(S i)) n) normalize + [>le_to_leb_true [2:@lt_to_le @(le_to_lt_to_lt … Hlt) //] + >not_eq_to_eqb_false // @lt_to_not_eq @(le_to_lt_to_lt … Hlt) // + |@(not_to_not … H) #Hn /2 by plus_minus/ + ] + |(not_le_to_leb_false (k+(S i)) n) normalize + [>(le_to_leb_true … Hle) >(not_eq_to_eqb_false … eqni) // + |@(not_to_not … nk) #H @le_plus_to_minus_r // + ] + ] + ] + |(cut (n < k+i)) [@(lt_to_le_to_lt ? i) /2 by not_le_to_lt/] #ltn + >not_le_to_leb_false [2: @lt_to_not_le @(transitive_lt …ltn) //] normalize + >not_le_to_leb_false [2: @lt_to_not_le //] normalize + >(not_le_to_leb_false … Hle) // + ] + |#M #N #HindM #HindN #i #HBC #HCc @eq_f2 [@HindM // |@HindN //] + |#ty #M #HindM #i #HBC #HCc @eq_f >plus_n_Sm >plus_n_Sm @HindM // + @is_closed_S // + |#ty #v #Hindv #i #HBC #HCc @eq_f + cut (∀m.mem ? m v → subst O D (subst O D m i B) (k+i) C = + subst O D (subst O D m (k+S i) C) i B) + [#m #Hmem @Hindv //] -Hindv elim v normalize [//] + #a #tl #Hind #H @eq_f2 [@H %1 // | @Hind #m #Hmem @H %2 //] + ] +qed. + + diff --git a/matita/matita/broken_lib/finite_lambda/typing.ma b/matita/matita/broken_lib/finite_lambda/typing.ma new file mode 100644 index 000000000..791e27d66 --- /dev/null +++ b/matita/matita/broken_lib/finite_lambda/typing.ma @@ -0,0 +1,229 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "finite_lambda/reduction.ma". + + +(****************************************************************) + +inductive TJ (O: Type[0]) (D:O → FinSet): list (FType O) → T O D → FType O → Prop ≝ + | tval: ∀G,o,a. TJ O D G (Val O D o a) (atom O o) + | trel: ∀G1,ty,G2,n. length ? G1 = n → TJ O D (G1@ty::G2) (Rel O D n) ty + | tapp: ∀G,M,N,ty1,ty2. TJ O D G M (arrow O ty1 ty2) → TJ O D G N ty1 → + TJ O D G (App O D M N) ty2 + | tlambda: ∀G,M,ty1,ty2. TJ O D (ty1::G) M ty2 → + TJ O D G (Lambda O D ty1 M) (arrow O ty1 ty2) + | tvec: ∀G,v,ty1,ty2. + (|v| = |enum (FinSet_of_FType O D ty1)|) → + (∀M. mem ? M v → TJ O D G M ty2) → + TJ O D G (Vec O D ty1 v) (arrow O ty1 ty2). + +lemma wt_to_T: ∀O,D,G,ty,a.TJ O D G (to_T O D ty a) ty. +#O #D #G #ty elim ty + [#o #a normalize @tval + |#ty1 #ty2 #Hind1 #Hind2 normalize * #v #Hv @tvec + [length_map >length_map // + |#M elim v + [normalize @False_ind |#a #v1 #Hind3 * [#eqM >eqM @Hind2 |@Hind3]] + ] + ] +qed. + +lemma inv_rel: ∀O,D,G,n,ty. + TJ O D G (Rel O D n) ty → ∃G1,G2.|G1|=n∧G=G1@ty::G2. +#O #D #G #n #ty #Hrel inversion Hrel + [#G1 #o #a #_ #H destruct + |#G1 #ty1 #G2 #n1 #H1 #H2 #H3 #H4 destruct %{G1} %{G2} /2/ + |#G1 #M0 #N #ty1 #ty2 #_ #_ #_ #_ #_ #H destruct + |#G1 #M0 #ty4 #ty5 #HM0 #_ #_ #H #H1 destruct + |#G1 #v #ty3 #ty4 #_ #_ #_ #_ #H destruct + ] +qed. + +lemma inv_tlambda: ∀O,D,G,M,ty1,ty2,ty3. + TJ O D G (Lambda O D ty1 M) (arrow O ty2 ty3) → + ty1 = ty2 ∧ TJ O D (ty2::G) M ty3. +#O #D #G #M #ty1 #ty2 #ty3 #Hlam inversion Hlam + [#G1 #o #a #_ #H destruct + |#G1 #ty #G2 #n #_ #_ #H destruct + |#G1 #M0 #N #ty1 #ty2 #_ #_ #_ #_ #_ #H destruct + |#G1 #M0 #ty4 #ty5 #HM0 #_ #_ #H #H1 destruct % // + |#G1 #v #ty3 #ty4 #_ #_ #_ #_ #H destruct + ] +qed. + +lemma inv_tvec: ∀O,D,G,v,ty1,ty2,ty3. + TJ O D G (Vec O D ty1 v) (arrow O ty2 ty3) → + (|v| = |enum (FinSet_of_FType O D ty1)|) ∧ + (∀M. mem ? M v → TJ O D G M ty3). +#O #D #G #v #ty1 #ty2 #ty3 #Hvec inversion Hvec + [#G #o #a #_ #H destruct + |#G1 #ty #G2 #n #_ #_ #H destruct + |#G1 #M0 #N #ty1 #ty2 #_ #_ #_ #_ #_ #H destruct + |#G1 #M0 #ty4 #ty5 #HM0 #_ #_ #H #H1 destruct + |#G1 #v1 #ty4 #ty5 #Hv #Hmem #_ #_ #H #H1 destruct % // @Hmem + ] +qed. + +(* could be generalized *) +lemma weak_rel: ∀O,D,G1,G2,ty1,ty2,n. length ? G1 < n → + TJ O D (G1@G2) (Rel O D n) ty1 → + TJ O D (G1@ty2::G2) (Rel O D (S n)) ty1. +#O #D #G1 #G2 #ty1 #ty2 #n #HG1 #Hrel lapply (inv_rel … Hrel) +* #G3 * #G4 * #H1 #H2 lapply (compare_append … H2) +* #G5 * + [* #H3 @False_ind >H3 in HG1; >length_append >H1 #H4 + @(absurd … H4) @le_to_not_lt // + |* #H3 #H4 >H4 >append_cons length_append >length_append

H3 >length_append normalize + >plus_n_Sm >associative_plus @eq_f // + ] +qed. + +lemma strength_rel: ∀O,D,G1,G2,ty1,ty2,n. length ? G1 < n → + TJ O D (G1@ty2::G2) (Rel O D n) ty1 → + TJ O D (G1@G2) (Rel O D (n-1)) ty1. +#O #D #G1 #G2 #ty1 #ty2 #n #HG1 #Hrel lapply (inv_rel … Hrel) +* #G3 * #G4 * #H1 #H2 lapply (compare_append … H2) +* #G5 * + [* #H3 @False_ind >H3 in HG1; >length_append >H1 #H4 + @(absurd … H4) @le_to_not_lt // + |lapply G5 -G5 * + [>append_nil normalize * #H3 #H4 destruct @False_ind @(absurd … HG1) + @le_to_not_lt // + |#ty3 #G5 * #H3 normalize #H4 destruct (H4) H3 >length_append >length_append normalize H1 >length_append // + |* cases G6 + [>append_nil normalize #H1 @False_ind + @(absurd ? Hlt) @le_to_not_lt H1 // + |#ty1 #G7 #H1 normalize #H2 destruct >associative_append @trel // + ] + ] + |#G #M #N #ty1 #ty2 #HM #HN #HindM #HindN #G1 #G2 #G3 + #Heq #Hc lapply (is_closed_app … Hc) -Hc * #HMc #HNc + @(tapp … (HindM … Heq HMc) (HindN … Heq HNc)) + |#G #M #ty1 #ty2 #HM #HindM #G1 #G2 #G3 #Heq #Hc + lapply (is_closed_lam … Hc) -Hc #HMc + @tlambda @(HindM (ty1::G1) G2) [>Heq // |@HMc] + |#G #v #ty1 #ty2 #Hlen #Hv #Hind #G1 #G2 #G3 #H1 #Hc @tvec + [>length_map // + |#M #Hmem @Hind // lapply (is_closed_vec … Hc) #Hvc @Hvc // + ] + ] +qed. + +lemma nth_spec: ∀A,a,d,l1,l2,n. |l1| = n → nth n A (l1@a::l2) d = a. +#A #a #d #l1 elim l1 normalize + [#l2 #n #Hn H cases (le_to_or_lt_eq … (leb_true_to_le … H)) + [#ltn >(not_eq_to_eqb_false … (lt_to_not_eq … ltn)) normalize + lapply (compare_append … HG) * #G3 * + [* #HG1 #HG2 @(strength_rel … tyN … ltn) HG in ltn; >length_append #ltn @False_ind + @(absurd … ltn) @le_to_not_lt >Hlen // + ] + |#HG21 >(eq_to_eqb_true … HG21) + cut (ty = tyN) + [<(nth_spec ? ty ty ? G2 … Hlen) >HG @nth_spec @HG21] #Hty >Hty + normalize H normalize lapply (compare_append … HG) * #G3 * + [* #H1 @False_ind @(absurd ? Hlen) @sym_not_eq @lt_to_not_eq >H1 + >length_append @(lt_to_le_to_lt n (|G21|)) // @not_le_to_lt + @(leb_false_to_not_le … H) + |cases G3 + [>append_nil * #H1 @False_ind @(absurd ? Hlen)

associative_append @trel // + ] + ] + ] + |#G #M #N #ty1 #ty2 #HM #HN #HindM #HindN #G1 #G2 #N0 #tyN0 #eqG + #HN0 #Hc normalize @(tapp … ty1) + [@(HindM … eqG HN0 Hc) |@(HindN … eqG HN0 Hc)] + |#G #M #ty1 #ty2 #HM #HindM #G1 #G2 #N0 #tyN0 #eqG + #HN0 #Hc normalize @(tlambda … ty1) @(HindM (ty1::G1) … HN0) // >eqG // + |#G #v #ty1 #ty2 #Hlen #Hv #Hind #G1 #G2 #N0 #tyN0 #eqG + #HN0 #Hc normalize @(tvec … ty1) + [>length_map @Hlen + |#M #Hmem lapply (mem_map ????? Hmem) * #a * -Hmem #Hmem #eqM HM1 @tlambda @Hind // + |#HM1 >HM1 @tvec // #N #HN lapply(mem_map ????? HN) + * #a * #mema #eqN H3 @tvec + [H2 >length_append >length_append @eq_f // + |#M2 #Hmem cases (mem_append ???? Hmem) -Hmem #Hmem + [@Hv >H2 @mem_append_l1 // + |cases Hmem + [#HM2 >HM2 -HM2 @(Hind N … H1) >H2 @mem_append_l2 %1 // + |-Hmem #Hmem @Hv >H2 @mem_append_l2 %2 // + ] + ] + ] + ] +qed. + diff --git a/matita/matita/broken_lib/reverse_complexity/big_O.ma b/matita/matita/broken_lib/reverse_complexity/big_O.ma new file mode 100644 index 000000000..833572ea2 --- /dev/null +++ b/matita/matita/broken_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/broken_lib/reverse_complexity/gap.ma b/matita/matita/broken_lib/reverse_complexity/gap.ma new file mode 100644 index 000000000..2c8b6e938 --- /dev/null +++ b/matita/matita/broken_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. + +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 new file mode 100644 index 000000000..9b9fecf94 --- /dev/null +++ b/matita/matita/broken_lib/reverse_complexity/speed_clean.ma @@ -0,0 +1,1067 @@ +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〉. + +(************************************* 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. + \ No newline at end of file diff --git a/matita/matita/broken_lib/reverse_complexity/speed_def.ma b/matita/matita/broken_lib/reverse_complexity/speed_def.ma new file mode 100644 index 000000000..6a0ec4a57 --- /dev/null +++ b/matita/matita/broken_lib/reverse_complexity/speed_def.ma @@ -0,0 +1,921 @@ +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〉. + +(************************************* 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. + \ No newline at end of file diff --git a/matita/matita/broken_lib/reverse_complexity/speed_new.ma b/matita/matita/broken_lib/reverse_complexity/speed_new.ma new file mode 100644 index 000000000..6a0ec4a57 --- /dev/null +++ b/matita/matita/broken_lib/reverse_complexity/speed_new.ma @@ -0,0 +1,921 @@ +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〉. + +(************************************* 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. + \ No newline at end of file diff --git a/matita/matita/broken_lib/reverse_complexity/toolkit.ma b/matita/matita/broken_lib/reverse_complexity/toolkit.ma new file mode 100644 index 000000000..11f988c79 --- /dev/null +++ b/matita/matita/broken_lib/reverse_complexity/toolkit.ma @@ -0,0 +1,987 @@ +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〉. + +(************************************* 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. + \ No newline at end of file diff --git a/matita/matita/broken_lib/turing/multi_to_mono/multi_to_mono.ma b/matita/matita/broken_lib/turing/multi_to_mono/multi_to_mono.ma new file mode 100644 index 000000000..0e9bce53e --- /dev/null +++ b/matita/matita/broken_lib/turing/multi_to_mono/multi_to_mono.ma @@ -0,0 +1,397 @@ +include "turing/auxiliary_machines1.ma". +include "turing/multi_to_mono/shift_trace_machines.ma". + +(******************************************************************************) +(* mtiL: complete move L for tape i. We reaching the left border of trace i, *) +(* add a blank if there is no more tape, then move the i-trace and finally *) +(* come back to the head position. *) +(******************************************************************************) + +(* we say that a tape is regular if for any trace after the first blank we + only have other blanks *) + +definition all_blanks_in ≝ λsig,l. + ∀x. mem ? x l → x = blank sig. + +definition regular_i ≝ λsig,n.λl:list (multi_sig sig n).λi. + all_blanks_in ? (after_blank ? (trace sig n i l)). + +definition regular_trace ≝ λsig,n,a.λls,rs:list (multi_sig sig n).λi. + Or (And (regular_i sig n (a::ls) i) (regular_i sig n rs i)) + (And (regular_i sig n ls i) (regular_i sig n (a::rs) i)). + +axiom regular_tail: ∀sig,n,l,i. + regular_i sig n l i → regular_i sig n (tail ? l) i. + +axiom regular_extend: ∀sig,n,l,i. + regular_i sig n l i → regular_i sig n (l@[all_blank sig n]) i. + +axiom all_blank_after_blank: ∀sig,n,l1,b,l2,i. + nth i ? (vec … b) (blank ?) = blank ? → + regular_i sig n (l1@b::l2) i → all_blanks_in ? (trace sig n i l2). + +lemma regular_trace_extl: ∀sig,n,a,ls,rs,i. + regular_trace sig n a ls rs i → + regular_trace sig n a (ls@[all_blank sig n]) rs i. +#sig #n #a #ls #rs #i * + [* #H1 #H2 % % // @(regular_extend … H1) + |* #H1 #H2 %2 % // @(regular_extend … H1) + ] +qed. + +lemma regular_cons_hd_rs: ∀sig,n.∀a:multi_sig sig n.∀ls,rs1,rs2,i. + regular_trace sig n a ls (rs1@rs2) i → + regular_trace sig n a ls (rs1@((hd ? rs2 (all_blank …))::(tail ? rs2))) i. +#sig #n #a #ls #rs1 #rs2 #i cases rs2 [2: #b #tl #H @H] +*[* #H1 >append_nil #H2 %1 % + [@H1 | whd in match (hd ???); @(regular_extend … rs1) //] + |* #H1 >append_nil #H2 %2 % + [@H1 | whd in match (hd ???); @(regular_extend … (a::rs1)) //] + ] +qed. + +lemma eq_trace_to_regular : ∀sig,n.∀a1,a2:multi_sig sig n.∀ls1,ls2,rs1,rs2,i. + nth i ? (vec … a1) (blank ?) = nth i ? (vec … a2) (blank ?) → + trace sig n i ls1 = trace sig n i ls2 → + trace sig n i rs1 = trace sig n i rs2 → + regular_trace sig n a1 ls1 rs1 i → + regular_trace sig n a2 ls2 rs2 i. +#sig #n #a1 #a2 #ls1 #ls2 #rs1 #rs2 #i #H1 #H2 #H3 #H4 +whd in match (regular_trace ??????); whd in match (regular_i ????); +whd in match (regular_i ?? rs2 ?); whd in match (regular_i ?? ls2 ?); +whd in match (regular_i ?? (a2::rs2) ?); whd in match (trace ????); +(Ht1a Hcur) in Ht2a; /2 by / + |#ls #a #rs #Htin #Hreg #Hreg2 -Ht1a cases (Ht1b … Htin) + [* #Hnb #Ht1 -Ht1b -Ht2a >Ht1 in Ht2b; >Htin #H + %{a} %{[ ]} %{ls} + %[%[%[%[%[@Hreg|@Hreg2]|@Hnb]|//]|//]|@H normalize % #H1 destruct (H1)] + |* + [(* we find the blank *) + * #ls1 * #b * #ls2 * * * #H1 #H2 #H3 #Ht1 + >Ht1 in Ht2b; #Hout -Ht1b + %{b} %{(a::ls1)} %{ls2} + %[%[%[%[%[>H1 in Hreg; #H @H + |#j #jneqi whd in match (hd ???); whd in match (tail ??); +

H1 //] + |@Hout normalize % normalize #H destruct (H) + ] + |* #b * #lss * * #H1 #H2 #Ht1 -Ht1b >Ht1 in Ht2a; + whd in match (left ??); whd in match (right ??); #Hout + %{(all_blank …)} %{(lss@[b])} %{[]} + %[%[%[%[%[

blank_all_blank //] + |

reverse_append >reverse_single @Hout normalize // + ] + ] + ] +qed. + +(******************************************************************************) + +definition shift_i_L ≝ λsig,n,i. + ncombf_r (multi_sig …) (shift_i sig n i) (all_blank sig n) · + mti sig n i · + extend ? (all_blank sig n). + +definition R_shift_i_L ≝ λsig,n,i,t1,t2. + (∀a,ls,rs. + t1 = midtape ? ls a rs → + ((∃rs1,b,rs2,a1,rss. + rs = rs1@b::rs2 ∧ + nth i ? (vec … b) (blank ?) = (blank ?) ∧ + (∀x. mem ? x rs1 → nth i ? (vec … x) (blank ?) ≠ (blank ?)) ∧ + shift_l sig n i (a::rs1) (a1::rss) ∧ + t2 = midtape (multi_sig sig n) ((reverse ? (a1::rss))@ls) b rs2) ∨ + (∃b,rss. + (∀x. mem ? x rs → nth i ? (vec … x) (blank ?) ≠ (blank ?)) ∧ + shift_l sig n i (a::rs) (rss@[b]) ∧ + t2 = midtape (multi_sig sig n) + ((reverse ? (rss@[b]))@ls) (all_blank sig n) [ ]))). + +definition R_shift_i_L_new ≝ λsig,n,i,t1,t2. + (∀a,ls,rs. + t1 = midtape ? ls a rs → + ∃rs1,b,rs2,rss. + b = hd ? rs2 (all_blank sig n) ∧ + nth i ? (vec … b) (blank ?) = (blank ?) ∧ + rs = rs1@rs2 ∧ + (∀x. mem ? x rs1 → nth i ? (vec … x) (blank ?) ≠ (blank ?)) ∧ + shift_l sig n i (a::rs1) rss ∧ + t2 = midtape (multi_sig sig n) ((reverse ? rss)@ls) b (tail ? rs2)). + +theorem sem_shift_i_L: ∀sig,n,i. shift_i_L sig n i ⊨ R_shift_i_L sig n i. +#sig #n #i +@(sem_seq_app ?????? + (sem_ncombf_r (multi_sig sig n) (shift_i sig n i)(all_blank sig n)) + (sem_seq ????? (ssem_mti sig n i) + (sem_extend ? (all_blank sig n)))) +#tin #tout * #t1 * * #Ht1a #Ht1b * #t2 * * #Ht2a #Ht2b * #Htout1 #Htout2 +#a #ls #rs cases rs + [#Htin %2 %{(shift_i sig n i a (all_blank sig n))} %{[ ]} + %[%[#x @False_ind | @daemon] + |lapply (Ht1a … Htin) -Ht1a -Ht1b #Ht1 + lapply (Ht2a … Ht1) -Ht2a -Ht2b #Ht2 >Ht2 in Htout1; + >Ht1 whd in match (left ??); whd in match (right ??); #Htout @Htout // + ] + |#a1 #rs1 #Htin + lapply (Ht1b … Htin) -Ht1a -Ht1b #Ht1 + lapply (Ht2b … Ht1) -Ht2a -Ht2b * + [(* a1 is blank *) * #H1 #H2 %1 + %{[ ]} %{a1} %{rs1} %{(shift_i sig n i a a1)} %{[ ]} + %[%[%[%[// |//] |#x @False_ind] | @daemon] + |>Htout2 [>H2 >reverse_single @Ht1 |>H2 >Ht1 normalize % #H destruct (H)] + ] + |* + [* #rs10 * #b * #rs2 * #rss * * * * #H1 #H2 #H3 #H4 + #Ht2 %1 + %{(a1::rs10)} %{b} %{rs2} %{(shift_i sig n i a a1)} %{rss} + %[%[%[%[>H1 //|//] |@H3] |@daemon ] + |>reverse_cons >associative_append + >H2 in Htout2; #Htout >Htout [@Ht2| >Ht2 normalize % #H destruct (H)] + ] + |* #b * #rss * * #H1 #H2 + #Ht2 %2 + %{(shift_i sig n i b (all_blank sig n))} %{(shift_i sig n i a a1::rss)} + %[%[@H1 |@daemon ] + |>Ht2 in Htout1; #Htout >Htout // + whd in match (left ??); whd in match (right ??); + >reverse_append >reverse_single >associative_append >reverse_cons + >associative_append // + ] + ] + ] + ] +qed. + +theorem sem_shift_i_L_new: ∀sig,n,i. + shift_i_L sig n i ⊨ R_shift_i_L_new sig n i. +#sig #n #i +@(Realize_to_Realize … (sem_shift_i_L sig n i)) +#t1 #t2 #H #a #ls #rs #Ht1 lapply (H a ls rs Ht1) * + [* #rs1 * #b * #rs2 * #a1 * #rss * * * * #H1 #H2 #H3 #H4 #Ht2 + %{rs1} %{b} %{(b::rs2)} %{(a1::rss)} + %[%[%[%[%[//|@H2]|@H1]|@H3]|@H4] | whd in match (tail ??); @Ht2] + |* #b * #rss * * #H1 #H2 #Ht2 + %{rs} %{(all_blank sig n)} %{[]} %{(rss@[b])} + %[%[%[%[%[//|@blank_all_blank]|//]|@H1]|@H2] | whd in match (tail ??); @Ht2] + ] +qed. + + +(******************************************************************************* +The following machine implements a full move of for a trace: we reach the left +border, shift the i-th trace and come back to the head position. *) + +(* this exclude the possibility that traces do not overlap: the head must +remain inside all traces *) + +definition mtiL ≝ λsig,n,i. + move_to_blank_L sig n i · + shift_i_L sig n i · + move_until ? L (no_head sig n). + +definition Rmtil ≝ λsig,n,i,t1,t2. + ∀ls,a,rs. + t1 = midtape (multi_sig sig n) ls a rs → + nth n ? (vec … a) (blank ?) = head ? → + (∀i.regular_trace sig n a ls rs i) → + (* next: we cannot be on rightof on trace i *) + (nth i ? (vec … a) (blank ?) = (blank ?) + → nth i ? (vec … (hd ? rs (all_blank …))) (blank ?) ≠ (blank ?)) → + no_head_in … ls → + no_head_in … rs → + (∃ls1,a1,rs1. + t2 = midtape (multi_sig …) ls1 a1 rs1 ∧ + (∀i.regular_trace … a1 ls1 rs1 i) ∧ + (∀j. j ≤ n → j ≠ i → to_blank_i ? n j (a1::ls1) = to_blank_i ? n j (a::ls)) ∧ + (∀j. j ≤ n → j ≠ i → to_blank_i ? n j rs1 = to_blank_i ? n j rs) ∧ + (to_blank_i ? n i ls1 = to_blank_i ? n i (a::ls)) ∧ + (to_blank_i ? n i (a1::rs1)) = to_blank_i ? n i rs). + +theorem sem_Rmtil: ∀sig,n,i. i < n → mtiL sig n i ⊨ Rmtil sig n i. +#sig #n #i #lt_in +@(sem_seq_app ?????? + (sem_move_to_blank_L … ) + (sem_seq ????? (sem_shift_i_L_new …) + (ssem_move_until_L ? (no_head sig n)))) +#tin #tout * #t1 * * #_ #Ht1 * #t2 * #Ht2 * #_ #Htout +(* we start looking into Rmitl *) +#ls #a #rs #Htin (* tin is a midtape *) +#Hhead #Hreg #no_rightof #Hnohead_ls #Hnohead_rs +cut (regular_i sig n (a::ls) i) + [cases (Hreg i) * // + cases (true_or_false (nth i ? (vec … a) (blank ?) == (blank ?))) #Htest + [#_ @daemon (* absurd, since hd rs non e' blank *) + |#H #_ @daemon]] #Hreg1 +lapply (Ht1 … Htin Hreg1 ?) [#j #_ @Hreg] -Ht1 -Htin +* #b * #ls1 * #ls2 * * * * * #reg_ls1_i #reg_ls1_j #Hno_blankb #Hhead #Hls1 #Ht1 +lapply (Ht2 … Ht1) -Ht2 -Ht1 +* #rs1 * #b0 * #rs2 * #rss * * * * * #Hb0 #Hb0blank #Hrs1 #Hrs1b #Hrss #Ht2 +(* we need to recover the position of the head of the emulated machine + that is the head of ls1. This is somewhere inside rs1 *) +cut (∃rs11. rs1 = (reverse ? ls1)@rs11) + [cut (ls1 = [ ] ∨ ∃aa,tlls1. ls1 = aa::tlls1) + [cases ls1 [%1 // | #aa #tlls1 %2 %{aa} %{tlls1} //]] * + [#H1ls1 %{rs1} >H1ls1 // + |* #aa * #tlls1 #H1ls1 >H1ls1 in Hrs1; + cut (aa = a) [>H1ls1 in Hls1; #H @(to_blank_hd … H)] #eqaa >eqaa + #Hrs1_aux cases (compare_append … (sym_eq … Hrs1_aux)) #l * + [* #H1 #H2 %{l} @H1 + |(* this is absurd : if l is empty, the case is as before. + if l is not empty then it must start with a blank, since it is the + first character in rs2. But in this case we would have a blank + inside ls1=a::tls1 that is absurd *) + @daemon + ]]] + * #rs11 #H1 +cut (rs = rs11@rs2) + [@(injective_append_l … (reverse … ls1)) >Hrs1 H2 >trace_append @mem_append_l2 + lapply Hb0 cases rs2 + [whd in match (hd ???); #H >H in H3; whd in match (no_head ???); + >all_blank_n normalize -H #H destruct (H); @False_ind + |#c #r #H4 %1 >H4 // + ] + |* + [(* we reach the head position *) + (* cut (trace sig n j (a1::ls20)=trace sig n j (ls1@b::ls2)) *) + * #ls10 * #a1 * #ls20 * * * #Hls20 #Ha1 #Hnh #Htout + cut (∀j.j ≠ i → + trace sig n j (reverse (multi_sig sig n) rs1@b::ls2) = + trace sig n j (ls10@a1::ls20)) + [#j #ineqj >append_cons trace_def reverse_map >map_append @eq_f @Hls20 ] + #Htracej + cut (trace sig n i (reverse (multi_sig sig n) (rs1@[b0])@ls2) = + trace sig n i (ls10@a1::ls20)) + [>trace_def Hcut + lapply (trace_shift … lt_in … Hrss) [//] whd in match (tail ??); #Htr reverse_map >map_append trace_def in ⊢ (%→?); trace_def in ⊢ (%→?); H1 in Htracei; >reverse_append >reverse_single >reverse_append + >reverse_reverse >associative_append >associative_append + @daemon + ] #H3 + cut (∀j. j ≠ i → + trace sig n j (reverse (multi_sig sig n) ls10@rs2) = trace sig n j rs) + [#j #jneqi @(injective_append_l … (trace sig n j (reverse ? ls1))) + >map_append >map_append >Hrs1 >H1 >associative_append + eqji (* by cases wether a1 is blank *) + @daemon + |#jneqi lapply (reg_ls1_j … jneqi) #H4 + >reverse_cons >associative_append >Hb0 @regular_cons_hd_rs + @(eq_trace_to_regular … H4) + [(proj2 … (H2 … jneqi)) >hd_trace % + |(proj2 … (H2 … jneqi)) >tail_trace % + |@sym_eq @Hrs_j // + ] + ]] + |#j #lejn #jneqi <(Hls1 … lejn) + >to_blank_i_def >to_blank_i_def @eq_f @sym_eq @(proj2 … (H2 j jneqi))] + |#j #lejn #jneqi >reverse_cons >associative_append >Hb0 + to_blank_i_def >to_blank_i_def @eq_f @Hrs_j //] + |<(Hls1 i) [2:@lt_to_le //] + lapply (all_blank_after_blank … reg_ls1_i) + [@(\P ?) @daemon] #allb_ls2 + whd in match (to_blank_i ????); <(proj2 … H3) + @daemon ] + |>reverse_cons >associative_append + cut (to_blank_i sig n i rs = to_blank_i sig n i (rs11@[b0])) [@daemon] + #Hcut >Hcut >(to_blank_i_chop … b0 (a1::reverse …ls10)) [2: @Hb0blank] + >to_blank_i_def >to_blank_i_def @eq_f + >trace_def >trace_def @injective_reverse >reverse_map >reverse_cons + >reverse_reverse >reverse_map >reverse_append >reverse_single @sym_eq + @(proj1 … H3) + ] + |(*we do not find the head: this is absurd *) + * #b1 * #lss * * #H2 @False_ind + cut (∀x0. mem ? x0 (trace sig n n (b0::reverse ? rss@ls2)) → x0 ≠ head ?) + [@daemon] -H2 #H2 + lapply (trace_shift_neq sig n i n … lt_in … Hrss) + [@lt_to_not_eq @lt_in | // ] + #H3 @(absurd + (nth n ? (vec … (hd ? (ls1@[b]) (all_blank sig n))) (blank ?) = head ?)) + [>Hhead // + |@H2 >trace_def %2 H3 >H1 >trace_def >reverse_map >reverse_cons >reverse_append + >reverse_reverse >associative_append (p_eqb3 ? y y …);//; -nlapply (refl ? (y ≟ x));ncases (y ≟ x) in ⊢ (???% → %);nnormalize;//; -#H1;napply p_eqb1;//; -nqed. -*) - -axiom swap_left : ∀N,x,y.swap N x y x = y. -(* -#N x y;nnormalize;nrewrite > (p_eqb3 ? x x …);//; -nqed. -*) - -axiom swap_other : ∀N,x,y,z.x ≠ z → y ≠ z → swap N x y z = z. -(* -#N x y z H1 H2;nnormalize;nrewrite > (p_eqb4 …); -##[nrewrite > (p_eqb4 …);//;@;ncases H2;/2/; -##|@;ncases H1;/2/ -##] -nqed. -*) - -axiom swap_inv : ∀N,x,y,z.swap N x y (swap N x y z) = z. -(* -#N x y z;nlapply (refl ? (x ≟ z));ncases (x ≟ z) in ⊢ (???% → ?);#H1 -##[nrewrite > (p_eqb1 … H1);nrewrite > (swap_left …);//; -##|nlapply (refl ? (y ≟ z));ncases (y ≟ z) in ⊢ (???% → ?);#H2 - ##[nrewrite > (p_eqb1 … H2);nrewrite > (swap_right …);//; - ##|nrewrite > (swap_other …) in ⊢ (??(????%)?);/2/; - nrewrite > (swap_other …);/2/; - ##] -##] -nqed. -*) - -axiom swap_fp : ∀N,x1,x2.finite_perm ? (swap N x1 x2). -(* -#N x1 x2;@ -##[@ - ##[nwhd;#xa xb;nnormalize;nlapply (refl ? (xa ≟ x1)); - ncases (xa ≟ x1) in ⊢ (???% → %);#H1 - ##[nrewrite > (p_eqb1 … H1);nlapply (refl ? (xb ≟ x1)); - ncases (xb ≟ x1) in ⊢ (???% → %);#H2 - ##[nrewrite > (p_eqb1 … H2);// - ##|nlapply (refl ? (xb ≟ x2)); - ncases (xb ≟ x2) in ⊢ (???% → %);#H3 - ##[nnormalize;#H4;nrewrite > H4 in H3; - #H3;nrewrite > H3 in H2;#H2;ndestruct (H2) - ##|nnormalize;#H4;nrewrite > H4 in H3; - nrewrite > (p_eqb3 …);//;#H5;ndestruct (H5) - ##] - ##] - ##|nlapply (refl ? (xa ≟ x2)); - ncases (xa ≟ x2) in ⊢ (???% → %);#H2 - ##[nrewrite > (p_eqb1 … H2);nlapply (refl ? (xb ≟ x1)); - ncases (xb ≟ x1) in ⊢ (???% → %);#H3 - ##[nnormalize;#H4;nrewrite > H4 in H3; - #H3;nrewrite > (p_eqb1 … H3);@ - ##|nnormalize;nlapply (refl ? (xb ≟ x2)); - ncases (xb ≟ x2) in ⊢ (???% → %);#H4 - ##[nrewrite > (p_eqb1 … H4);// - ##|nnormalize;#H5;nrewrite > H5 in H3; - nrewrite > (p_eqb3 …);//;#H6;ndestruct (H6); - ##] - ##] - ##|nnormalize;nlapply (refl ? (xb ≟ x1)); - ncases (xb ≟ x1) in ⊢ (???% → %);#H3 - ##[nnormalize;#H4;nrewrite > H4 in H2;nrewrite > (p_eqb3 …);//; - #H5;ndestruct (H5) - ##|nlapply (refl ? (xb ≟ x2)); - ncases (xb ≟ x2) in ⊢ (???% → %);#H4 - ##[nnormalize;#H5;nrewrite > H5 in H1;nrewrite > (p_eqb3 …);//; - #H6;ndestruct (H6) - ##|nnormalize;// - ##] - ##] - ##] - ##] - ##|nwhd;#z;nnormalize;nlapply (refl ? (z ≟ x1)); - ncases (z ≟ x1) in ⊢ (???% → %);#H1 - ##[nlapply (refl ? (z ≟ x2)); - ncases (z ≟ x2) in ⊢ (???% → %);#H2 - ##[@ z;nrewrite > H1;nrewrite > H2;napply p_eqb1;// - ##|@ x2;nrewrite > (p_eqb4 …); - ##[nrewrite > (p_eqb3 …);//; - nnormalize;napply p_eqb1;// - ##|nrewrite < (p_eqb1 … H1);@;#H3;nrewrite > H3 in H2; - nrewrite > (p_eqb3 …);//;#H2;ndestruct (H2) - ##] - ##] - ##|nlapply (refl ? (z ≟ x2)); - ncases (z ≟ x2) in ⊢ (???% → %);#H2 - ##[@ x1;nrewrite > (p_eqb3 …);//; - napply p_eqb1;nnormalize;// - ##|@ z;nrewrite > H1;nrewrite > H2;@; - ##] - ##] - ##] -##|@ [x1;x2];#x0 H1;nrewrite > (swap_other …) - ##[@ - ##|@;#H2;nrewrite > H2 in H1;*;#H3;napply H3;/2/; - ##|@;#H2;nrewrite > H2 in H1;*;#H3;napply H3;//; - ##] -##] -nqed. -*) - -axiom inj_swap : ∀N,u,v.injective ?? (swap N u v). -(* -#N u v;ncases (swap_fp N u v);*;#H1 H2 H3;//; -nqed. -*) - -axiom surj_swap : ∀N,u,v.surjective ?? (swap N u v). -(* -#N u v;ncases (swap_fp N u v);*;#H1 H2 H3;//; -nqed. -*) - -axiom finite_swap : ∀N,u,v.∃l.∀x.x ∉ l → swap N u v x = x. -(* -#N u v;ncases (swap_fp N u v);*;#H1 H2 H3;//; -nqed. -*) - -(* swaps two lists of parameters -definition Pi_swap_list : ∀xl,xl':list X.X → X ≝ - λxl,xl',x.foldr2 ??? (λu,v,r.swap ? u v r) x xl xl'. - -nlemma fp_swap_list : - ∀xl,xl'.finite_perm ? (Pi_swap_list xl xl'). -#xl xl';@ -##[@; - ##[ngeneralize in match xl';nelim xl - ##[nnormalize;//; - ##|#x0 xl0;#IH xl'';nelim xl'' - ##[nnormalize;// - ##|#x1 xl1 IH1 y0 y1;nchange in ⊢ (??%% → ?) with (swap ????); - #H1;nlapply (inj_swap … H1);#H2; - nlapply (IH … H2);// - ##] - ##] - ##|ngeneralize in match xl';nelim xl - ##[nnormalize;#_;#z;@z;@ - ##|#x' xl0 IH xl'';nelim xl'' - ##[nnormalize;#z;@z;@ - ##|#x1 xl1 IH1 z; - nchange in ⊢ (??(λ_.???%)) with (swap ????); - ncases (surj_swap X x' x1 z);#x2 H1; - ncases (IH xl1 x2);#x3 H2;@ x3; - nrewrite < H2;napply H1 - ##] - ##] - ##] -##|ngeneralize in match xl';nelim xl - ##[#;@ [];#;@ - ##|#x0 xl0 IH xl'';ncases xl'' - ##[@ [];#;@ - ##|#x1 xl1;ncases (IH xl1);#xl2 H1; - ncases (finite_swap X x0 x1);#xl3 H2; - @ (xl2@xl3);#x2 H3; - nchange in ⊢ (??%?) with (swap ????); - nrewrite > (H1 …); - ##[nrewrite > (H2 …);//;@;#H4;ncases H3;#H5;napply H5; - napply in_list_to_in_list_append_r;// - ##|@;#H4;ncases H3;#H5;napply H5; - napply in_list_to_in_list_append_l;// - ##] - ##] - ##] -##] -nqed. - -(* the 'reverse' swap of lists of parameters - composing Pi_swap_list and Pi_swap_list_r yields the identity function - (see the Pi_swap_list_inv lemma) *) -ndefinition Pi_swap_list_r : ∀xl,xl':list X. Pi ≝ - λxl,xl',x.foldl2 ??? (λr,u,v.swap ? u v r ) x xl xl'. - -nlemma fp_swap_list_r : - ∀xl,xl'.finite_perm ? (Pi_swap_list_r xl xl'). -#xl xl';@ -##[@; - ##[ngeneralize in match xl';nelim xl - ##[nnormalize;//; - ##|#x0 xl0;#IH xl'';nelim xl'' - ##[nnormalize;// - ##|#x1 xl1 IH1 y0 y1;nwhd in ⊢ (??%% → ?); - #H1;nlapply (IH … H1);#H2; - napply (inj_swap … H2); - ##] - ##] - ##|ngeneralize in match xl';nelim xl - ##[nnormalize;#_;#z;@z;@ - ##|#x' xl0 IH xl'';nelim xl'' - ##[nnormalize;#z;@z;@ - ##|#x1 xl1 IH1 z;nwhd in ⊢ (??(λ_.???%)); - ncases (IH xl1 z);#x2 H1; - ncases (surj_swap X x' x1 x2);#x3 H2; - @ x3;nrewrite < H2;napply H1; - ##] - ##] - ##] -##|ngeneralize in match xl';nelim xl - ##[#;@ [];#;@ - ##|#x0 xl0 IH xl'';ncases xl'' - ##[@ [];#;@ - ##|#x1 xl1; - ncases (IH xl1);#xl2 H1; - ncases (finite_swap X x0 x1);#xl3 H2; - @ (xl2@xl3);#x2 H3;nwhd in ⊢ (??%?); - nrewrite > (H2 …); - ##[nrewrite > (H1 …);//;@;#H4;ncases H3;#H5;napply H5; - napply in_list_to_in_list_append_l;// - ##|@;#H4;ncases H3;#H5;napply H5; - napply in_list_to_in_list_append_r;// - ##] - ##] - ##] -##] -nqed. - -nlemma Pi_swap_list_inv : - ∀xl1,xl2,x. - Pi_swap_list xl1 xl2 (Pi_swap_list_r xl1 xl2 x) = x. -#xl;nelim xl -##[#;@ -##|#x1 xl1 IH xl';ncases xl' - ##[#;@ - ##|#x2 xl2;#x; - nchange in ⊢ (??%?) with - (swap ??? (Pi_swap_list ?? - (Pi_swap_list_r ?? (swap ????)))); - nrewrite > (IH xl2 ?);napply swap_inv; - ##] -##] -nqed. - -nlemma Pi_swap_list_fresh : - ∀x,xl1,xl2.x ∉ xl1 → x ∉ xl2 → Pi_swap_list xl1 xl2 x = x. -#x xl1;nelim xl1 -##[#;@ -##|#x3 xl3 IH xl2' H1;ncases xl2' - ##[#;@ - ##|#x4 xl4 H2;ncut (x ∉ xl3 ∧ x ∉ xl4); - ##[@ - ##[@;#H3;ncases H1;#H4;napply H4;/2/; - ##|@;#H3;ncases H2;#H4;napply H4;/2/ - ##] - ##] *; #H1' H2'; - nchange in ⊢ (??%?) with (swap ????); - nrewrite > (swap_other …) - ##[napply IH;//; - ##|nchange in ⊢ (?(???%)) with (Pi_swap_list ???); - nrewrite > (IH …);//;@;#H3;ncases H2;#H4;napply H4;//; - ##|nchange in ⊢ (?(???%)) with (Pi_swap_list ???); - nrewrite > (IH …);//;@;#H3;ncases H1;#H4;napply H4;// - ##] - ##] -##] -nqed. -*) \ No newline at end of file diff --git a/matita/matita/lib/binding/ln.ma b/matita/matita/lib/binding/ln.ma deleted file mode 100644 index ca7f574ed..000000000 --- a/matita/matita/lib/binding/ln.ma +++ /dev/null @@ -1,716 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basics/lists/list.ma". -include "basics/deqsets.ma". -include "binding/names.ma". -include "binding/fp.ma". - -axiom alpha : Nset. -notation "𝔸" non associative with precedence 90 for @{'alphabet}. -interpretation "set of names" 'alphabet = alpha. - -inductive tp : Type[0] ≝ -| top : tp -| arr : tp → tp → tp. -inductive pretm : Type[0] ≝ -| var : nat → pretm -| par : 𝔸 → pretm -| abs : tp → pretm → pretm -| app : pretm → pretm → pretm. - -let rec Nth T n (l:list T) on n ≝ - match l with - [ nil ⇒ None ? - | cons hd tl ⇒ match n with - [ O ⇒ Some ? hd - | S n0 ⇒ Nth T n0 tl ] ]. - -let rec vclose_tm_aux u x k ≝ match u with - [ var n ⇒ if (leb k n) then var (S n) else u - | par x0 ⇒ if (x0 == x) then (var k) else u - | app v1 v2 ⇒ app (vclose_tm_aux v1 x k) (vclose_tm_aux v2 x k) - | abs s v ⇒ abs s (vclose_tm_aux v x (S k)) ]. -definition vclose_tm ≝ λu,x.vclose_tm_aux u x O. - -definition vopen_var ≝ λn,x,k.match eqb n k with - [ true ⇒ par x - | false ⇒ match leb n k with - [ true ⇒ var n - | false ⇒ var (pred n) ] ]. - -let rec vopen_tm_aux u x k ≝ match u with - [ var n ⇒ vopen_var n x k - | par x0 ⇒ u - | app v1 v2 ⇒ app (vopen_tm_aux v1 x k) (vopen_tm_aux v2 x k) - | abs s v ⇒ abs s (vopen_tm_aux v x (S k)) ]. -definition vopen_tm ≝ λu,x.vopen_tm_aux u x O. - -let rec FV u ≝ match u with - [ par x ⇒ [x] - | app v1 v2 ⇒ FV v1@FV v2 - | abs s v ⇒ FV v - | _ ⇒ [ ] ]. - -definition lam ≝ λx,s,u.abs s (vclose_tm u x). - -let rec Pi_map_tm p u on u ≝ match u with -[ par x ⇒ par (p x) -| var _ ⇒ u -| app v1 v2 ⇒ app (Pi_map_tm p v1) (Pi_map_tm p v2) -| abs s v ⇒ abs s (Pi_map_tm p v) ]. - -interpretation "permutation of tm" 'middot p x = (Pi_map_tm p x). - -notation "hvbox(u⌈x⌉)" - with precedence 45 - for @{ 'open $u $x }. - -(* -notation "hvbox(u⌈x⌉)" - with precedence 45 - for @{ 'open $u $x }. -notation "❴ u ❵ x" non associative with precedence 90 for @{ 'open $u $x }. -*) -interpretation "ln term variable open" 'open u x = (vopen_tm u x). -notation < "hvbox(ν x break . u)" - with precedence 20 -for @{'nu $x $u }. -notation > "ν list1 x sep , . term 19 u" with precedence 20 - for ${ fold right @{$u} rec acc @{'nu $x $acc)} }. -interpretation "ln term variable close" 'nu x u = (vclose_tm u x). - -let rec tm_height u ≝ match u with -[ app v1 v2 ⇒ S (max (tm_height v1) (tm_height v2)) -| abs s v ⇒ S (tm_height v) -| _ ⇒ O ]. - -theorem le_n_O_rect_Type0 : ∀n:nat. n ≤ O → ∀P: nat →Type[0]. P O → P n. -#n (cases n) // #a #abs cases (?:False) /2/ qed. - -theorem nat_rect_Type0_1 : ∀n:nat.∀P:nat → Type[0]. -(∀m.(∀p. p < m → P p) → P m) → P n. -#n #P #H -cut (∀q:nat. q ≤ n → P q) /2/ -(elim n) - [#q #HleO (* applica male *) - @(le_n_O_rect_Type0 ? HleO) - @H #p #ltpO cases (?:False) /2/ (* 3 *) - |#p #Hind #q #HleS - @H #a #lta @Hind @le_S_S_to_le /2/ - ] -qed. - -lemma leb_false_to_lt : ∀n,m. leb n m = false → m < n. -#n elim n -[ #m normalize #H destruct(H) -| #n0 #IH * // #m normalize #H @le_S_S @IH // ] -qed. - -lemma nominal_eta_aux : ∀x,u.x ∉ FV u → ∀k.vclose_tm_aux (vopen_tm_aux u x k) x k = u. -#x #u elim u -[ #n #_ #k normalize cases (decidable_eq_nat n k) #Hnk - [ >Hnk >eqb_n_n normalize >(\b ?) // - | >(not_eq_to_eqb_false … Hnk) normalize cases (true_or_false (leb n k)) #Hleb - [ >Hleb normalize >(?:leb k n = false) // - @lt_to_leb_false @not_eq_to_le_to_lt /2/ - | >Hleb normalize >(?:leb k (pred n) = true) normalize - [ cases (leb_false_to_lt … Hleb) // - | @le_to_leb_true cases (leb_false_to_lt … Hleb) normalize /2/ ] ] ] -| #y normalize #Hy >(\bf ?) // @(not_to_not … Hy) // -| #s #v #IH normalize #Hv #k >IH // @Hv -| #v1 #v2 #IH1 #IH2 normalize #Hv1v2 #k - >IH1 [ >IH2 // | @(not_to_not … Hv1v2) @in_list_to_in_list_append_l ] - @(not_to_not … Hv1v2) @in_list_to_in_list_append_r ] -qed. - -corollary nominal_eta : ∀x,u.x ∉ FV u → (νx.u⌈x⌉) = u. -#x #u #Hu @nominal_eta_aux // -qed. - -lemma eq_height_vopen_aux : ∀v,x,k.tm_height (vopen_tm_aux v x k) = tm_height v. -#v #x elim v -[ #n #k normalize cases (eqb n k) // cases (leb n k) // -| #u #k % -| #s #u #IH #k normalize >IH % -| #u1 #u2 #IH1 #IH2 #k normalize >IH1 >IH2 % ] -qed. - -corollary eq_height_vopen : ∀v,x.tm_height (v⌈x⌉) = tm_height v. -#v #x @eq_height_vopen_aux -qed. - -theorem pretm_ind_plus_aux : - ∀P:pretm → Type[0]. - (∀x:𝔸.P (par x)) → - (∀n:ℕ.P (var n)) → - (∀v1,v2. P v1 → P v2 → P (app v1 v2)) → - ∀C:list 𝔸. - (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) → - ∀n,u.tm_height u ≤ n → P u. -#P #Hpar #Hvar #Happ #C #Hlam #n change with ((λn.?) n); @(nat_rect_Type0_1 n ??) -#m cases m -[ #_ * /2/ - [ normalize #s #v #Hfalse cases (?:False) cases (not_le_Sn_O (tm_height v)) /2/ - | #v1 #v2 whd in ⊢ (?%?→?); #Hfalse cases (?:False) cases (not_le_Sn_O (max ??)) - [ #H @H @Hfalse|*:skip] ] ] --m #m #IH * /2/ -[ #s #v whd in ⊢ (?%?→?); #Hv - lapply (p_fresh … (C@FV v)) letin y ≝ (N_fresh … (C@FV v)) #Hy - >(?:abs s v = lam y s (v⌈y⌉)) - [| whd in ⊢ (???%); >nominal_eta // @(not_to_not … Hy) @in_list_to_in_list_append_r ] - @Hlam - [ @(not_to_not … Hy) @in_list_to_in_list_append_r - | @(not_to_not … Hy) @in_list_to_in_list_append_l ] - @IH [| @Hv | >eq_height_vopen % ] -| #v1 #v2 whd in ⊢ (?%?→?); #Hv @Happ - [ @IH [| @Hv | @le_max_1 ] | @IH [| @Hv | @le_max_2 ] ] ] -qed. - -corollary pretm_ind_plus : - ∀P:pretm → Type[0]. - (∀x:𝔸.P (par x)) → - (∀n:ℕ.P (var n)) → - (∀v1,v2. P v1 → P v2 → P (app v1 v2)) → - ∀C:list 𝔸. - (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) → - ∀u.P u. -#P #Hpar #Hvar #Happ #C #Hlam #u @pretm_ind_plus_aux /2/ -qed. - -(* maps a permutation to a list of terms *) -definition Pi_map_list : (𝔸 → 𝔸) → list 𝔸 → list 𝔸 ≝ map 𝔸 𝔸 . - -(* interpretation "permutation of name list" 'middot p x = (Pi_map_list p x).*) - -(* -inductive tm : pretm → Prop ≝ -| tm_par : ∀x:𝔸.tm (par x) -| tm_app : ∀u,v.tm u → tm v → tm (app u v) -| tm_lam : ∀x,s,u.tm u → tm (lam x s u). - -inductive ctx_aux : nat → pretm → Prop ≝ -| ctx_var : ∀n,k.n < k → ctx_aux k (var n) -| ctx_par : ∀x,k.ctx_aux k (par x) -| ctx_app : ∀u,v,k.ctx_aux k u → ctx_aux k v → ctx_aux k (app u v) -(* è sostituibile da ctx_lam ? *) -| ctx_abs : ∀s,u.ctx_aux (S k) u → ctx_aux k (abs s u). -*) - -inductive tm_or_ctx (k:nat) : pretm → Type[0] ≝ -| toc_var : ∀n.n < k → tm_or_ctx k (var n) -| toc_par : ∀x.tm_or_ctx k (par x) -| toc_app : ∀u,v.tm_or_ctx k u → tm_or_ctx k v → tm_or_ctx k (app u v) -| toc_lam : ∀x,s,u.tm_or_ctx k u → tm_or_ctx k (lam x s u). - -definition tm ≝ λt.tm_or_ctx O t. -definition ctx ≝ λt.tm_or_ctx 1 t. - -definition check_tm ≝ λu,k. - pretm_ind_plus ? - (λ_.true) - (λn.leb (S n) k) - (λv1,v2,rv1,rv2.rv1 ∧ rv2) - [] (λx,s,v,px,pC,rv.rv) - u. - -axiom pretm_ind_plus_app : ∀P,u,v,C,H1,H2,H3,H4. - pretm_ind_plus P H1 H2 H3 C H4 (app u v) = - H3 u v (pretm_ind_plus P H1 H2 H3 C H4 u) (pretm_ind_plus P H1 H2 H3 C H4 v). - -axiom pretm_ind_plus_lam : ∀P,x,s,u,C,px,pC,H1,H2,H3,H4. - pretm_ind_plus P H1 H2 H3 C H4 (lam x s (u⌈x⌉)) = - H4 x s u px pC (pretm_ind_plus P H1 H2 H3 C H4 (u⌈x⌉)). - -record TM : Type[0] ≝ { - pretm_of_TM :> pretm; - tm_of_TM : check_tm pretm_of_TM O = true -}. - -record CTX : Type[0] ≝ { - pretm_of_CTX :> pretm; - ctx_of_CTX : check_tm pretm_of_CTX 1 = true -}. - -inductive tm2 : pretm → Type[0] ≝ -| tm_par : ∀x.tm2 (par x) -| tm_app : ∀u,v.tm2 u → tm2 v → tm2 (app u v) -| tm_lam : ∀x,s,u.x ∉ FV u → (∀y.y ∉ FV u → tm2 (u⌈y⌉)) → tm2 (lam x s (u⌈x⌉)). - -(* -inductive tm' : pretm → Prop ≝ -| tm_par : ∀x.tm' (par x) -| tm_app : ∀u,v.tm' u → tm' v → tm' (app u v) -| tm_lam : ∀x,s,u,C.x ∉ FV u → x ∉ C → (∀y.y ∉ FV u → tm' (❴u❵y)) → tm' (lam x s (❴u❵x)). -*) - -lemma pi_vclose_tm : - ∀z1,z2,x,u.swap 𝔸 z1 z2·(νx.u) = (ν swap ? z1 z2 x.swap 𝔸 z1 z2 · u). -#z1 #z2 #x #u -change with (vclose_tm_aux ???) in match (vclose_tm ??); -change with (vclose_tm_aux ???) in ⊢ (???%); lapply O elim u normalize // -[ #n #k cases (leb k n) normalize % -| #x0 #k cases (true_or_false (x0==z1)) #H1 >H1 normalize - [ cases (true_or_false (x0==x)) #H2 >H2 normalize - [ <(\P H2) >H1 normalize >(\b (refl ? z2)) % - | >H1 normalize cases (true_or_false (x==z1)) #H3 >H3 normalize - [ >(\P H3) in H2; >H1 #Hfalse destruct (Hfalse) - | cases (true_or_false (x==z2)) #H4 >H4 normalize - [ cases (true_or_false (z2==z1)) #H5 >H5 normalize // - >(\P H5) in H4; >H3 #Hfalse destruct (Hfalse) - | >(\bf ?) // @sym_not_eq @(\Pf H4) ] - ] - ] - | cases (true_or_false (x0==x)) #H2 >H2 normalize - [ <(\P H2) >H1 normalize >(\b (refl ??)) % - | >H1 normalize cases (true_or_false (x==z1)) #H3 >H3 normalize - [ cases (true_or_false (x0==z2)) #H4 >H4 normalize - [ cases (true_or_false (z1==z2)) #H5 >H5 normalize // - <(\P H5) in H4; <(\P H3) >H2 #Hfalse destruct (Hfalse) - | >H4 % ] - | cases (true_or_false (x0==z2)) #H4 >H4 normalize - [ cases (true_or_false (x==z2)) #H5 >H5 normalize - [ <(\P H5) in H4; >H2 #Hfalse destruct (Hfalse) - | >(\bf ?) // @sym_not_eq @(\Pf H3) ] - | cases (true_or_false (x==z2)) #H5 >H5 normalize - [ >H1 % - | >H2 % ] - ] - ] - ] - ] -] -qed. - -lemma pi_vopen_tm : - ∀z1,z2,x,u.swap 𝔸 z1 z2·(u⌈x⌉) = (swap 𝔸 z1 z2 · u⌈swap 𝔸 z1 z2 x⌉). -#z1 #z2 #x #u -change with (vopen_tm_aux ???) in match (vopen_tm ??); -change with (vopen_tm_aux ???) in ⊢ (???%); lapply O elim u normalize // -#n #k cases (true_or_false (eqb n k)) #H1 >H1 normalize // -cases (true_or_false (leb n k)) #H2 >H2 normalize // -qed. - -lemma pi_lam : - ∀z1,z2,x,s,u.swap 𝔸 z1 z2 · lam x s u = lam (swap 𝔸 z1 z2 x) s (swap 𝔸 z1 z2 · u). -#z1 #z2 #x #s #u whd in ⊢ (???%); <(pi_vclose_tm …) % -qed. - -lemma eqv_FV : ∀z1,z2,u.FV (swap 𝔸 z1 z2 · u) = Pi_map_list (swap 𝔸 z1 z2) (FV u). -#z1 #z2 #u elim u // -[ #s #v normalize // -| #v1 #v2 normalize /2/ ] -qed. - -lemma swap_inv_tm : ∀z1,z2,u.swap 𝔸 z1 z2 · (swap 𝔸 z1 z2 · u) = u. -#z1 #z2 #u elim u [1,3,4:normalize //] -#x whd in ⊢ (??%?); >swap_inv % -qed. - -lemma eqv_in_list : ∀x,l,z1,z2.x ∈ l → swap 𝔸 z1 z2 x ∈ Pi_map_list (swap 𝔸 z1 z2) l. -#x #l #z1 #z2 #Hin elim Hin -[ #x0 #l0 % -| #x1 #x2 #l0 #Hin #IH %2 @IH ] -qed. - -lemma eqv_tm2 : ∀u.tm2 u → ∀z1,z2.tm2 ((swap ? z1 z2)·u). -#u #Hu #z1 #z2 letin p ≝ (swap ? z1 z2) elim Hu /2/ -#x #s #v #Hx #Hv #IH >pi_lam >pi_vopen_tm %3 -[ @(not_to_not … Hx) -Hx #Hx - <(swap_inv ? z1 z2 x) <(swap_inv_tm z1 z2 v) >eqv_FV @eqv_in_list // -| #y #Hy <(swap_inv ? z1 z2 y) - eqv_FV @eqv_in_list // -] -qed. - -lemma vclose_vopen_aux : ∀x,u,k.vopen_tm_aux (vclose_tm_aux u x k) x k = u. -#x #u elim u normalize // -[ #n #k cases (true_or_false (leb k n)) #H >H whd in ⊢ (??%?); - [ cases (true_or_false (eqb (S n) k)) #H1 >H1 - [ <(eqb_true_to_eq … H1) in H; #H lapply (leb_true_to_le … H) -H #H - cases (le_to_not_lt … H) -H #H cases (H ?) % - | whd in ⊢ (??%?); >lt_to_leb_false // @le_S_S /2/ ] - | cases (true_or_false (eqb n k)) #H1 >H1 normalize - [ >(eqb_true_to_eq … H1) in H; #H lapply (leb_false_to_not_le … H) -H - * #H cases (H ?) % - | >le_to_leb_true // @not_lt_to_le % #H2 >le_to_leb_true in H; - [ #H destruct (H) | /2/ ] - ] - ] -| #x0 #k cases (true_or_false (x0==x)) #H1 >H1 normalize // >(\P H1) >eqb_n_n % ] -qed. - -lemma vclose_vopen : ∀x,u.((νx.u)⌈x⌉) = u. #x #u @vclose_vopen_aux -qed. - -(* -theorem tm_to_tm : ∀t.tm' t → tm t. -#t #H elim H -*) - -lemma in_list_singleton : ∀T.∀t1,t2:T.t1 ∈ [t2] → t1 = t2. -#T #t1 #t2 #H @(in_list_inv_ind ??? H) /2/ -qed. - -lemma fresh_vclose_tm_aux : ∀u,x,k.x ∉ FV (vclose_tm_aux u x k). -#u #x elim u // -[ #n #k normalize cases (leb k n) normalize // -| #x0 #k normalize cases (true_or_false (x0==x)) #H >H normalize // - lapply (\Pf H) @not_to_not #Hin >(in_list_singleton ??? Hin) % -| #v1 #v2 #IH1 #IH2 #k normalize % #Hin cases (in_list_append_to_or_in_list ???? Hin) /2/ ] -qed. - -lemma fresh_vclose_tm : ∀u,x.x ∉ FV (νx.u). // -qed. - -lemma check_tm_true_to_toc : ∀u,k.check_tm u k = true → tm_or_ctx k u. -#u @(pretm_ind_plus ???? [ ] ? u) -[ #x #k #_ %2 -| #n #k change with (leb (S n) k) in ⊢ (??%?→?); #H % @leb_true_to_le // -| #v1 #v2 #rv1 #rv2 #k change with (pretm_ind_plus ???????) in ⊢ (??%?→?); - >pretm_ind_plus_app #H cases (andb_true ?? H) -H #Hv1 #Hv2 %3 - [ @rv1 @Hv1 | @rv2 @Hv2 ] -| #x #s #v #Hx #_ #rv #k change with (pretm_ind_plus ???????) in ⊢ (??%?→?); - >pretm_ind_plus_lam // #Hv %4 @rv @Hv ] -qed. - -lemma toc_to_check_tm_true : ∀u,k.tm_or_ctx k u → check_tm u k = true. -#u #k #Hu elim Hu // -[ #n #Hn change with (leb (S n) k) in ⊢ (??%?); @le_to_leb_true @Hn -| #v1 #v2 #Hv1 #Hv2 #IH1 #IH2 change with (pretm_ind_plus ???????) in ⊢ (??%?); - >pretm_ind_plus_app change with (check_tm v1 k ∧ check_tm v2 k) in ⊢ (??%?); /2/ -| #x #s #v #Hv #IH <(vclose_vopen x v) change with (pretm_ind_plus ???????) in ⊢ (??%?); - >pretm_ind_plus_lam [| // | @fresh_vclose_tm ] >(vclose_vopen x v) @IH ] -qed. - -lemma fresh_swap_tm : ∀z1,z2,u.z1 ∉ FV u → z2 ∉ FV u → swap 𝔸 z1 z2 · u = u. -#z1 #z2 #u elim u -[2: normalize in ⊢ (?→%→%→?); #x #Hz1 #Hz2 whd in ⊢ (??%?); >swap_other // - [ @(not_to_not … Hz2) | @(not_to_not … Hz1) ] // -|1: // -| #s #v #IH normalize #Hz1 #Hz2 >IH // [@Hz2|@Hz1] -| #v1 #v2 #IH1 #IH2 normalize #Hz1 #Hz2 - >IH1 [| @(not_to_not … Hz2) @in_list_to_in_list_append_l | @(not_to_not … Hz1) @in_list_to_in_list_append_l ] - >IH2 // [@(not_to_not … Hz2) @in_list_to_in_list_append_r | @(not_to_not … Hz1) @in_list_to_in_list_append_r ] -] -qed. - -theorem tm_to_tm2 : ∀u.tm u → tm2 u. -#t #Ht elim Ht -[ #n #Hn cases (not_le_Sn_O n) #Hfalse cases (Hfalse Hn) -| @tm_par -| #u #v #Hu #Hv @tm_app -| #x #s #u #Hu #IHu <(vclose_vopen x u) @tm_lam - [ @fresh_vclose_tm - | #y #Hy <(fresh_swap_tm x y (νx.u)) /2/ @fresh_vclose_tm ] -] -qed. - -theorem tm2_to_tm : ∀u.tm2 u → tm u. -#u #pu elim pu /2/ #x #s #v #Hx #Hv #IH %4 @IH // -qed. - -(* define PAR APP LAM *) -definition PAR ≝ λx.mk_TM (par x) ?. // qed. -definition APP ≝ λu,v:TM.mk_TM (app u v) ?. -change with (pretm_ind_plus ???????) in match (check_tm ??); >pretm_ind_plus_app -change with (check_tm ??) in match (pretm_ind_plus ???????); change with (check_tm ??) in match (pretm_ind_plus ???????) in ⊢ (??(??%)?); -@andb_elim >(tm_of_TM u) >(tm_of_TM v) % qed. -definition LAM ≝ λx,s.λu:TM.mk_TM (lam x s u) ?. -change with (pretm_ind_plus ???????) in match (check_tm ??); <(vclose_vopen x u) ->pretm_ind_plus_lam [| // | @fresh_vclose_tm ] -change with (check_tm ??) in match (pretm_ind_plus ???????); >vclose_vopen -@(tm_of_TM u) qed. - -axiom vopen_tm_down : ∀u,x,k.tm_or_ctx (S k) u → tm_or_ctx k (u⌈x⌉). -(* needs true_plus_false - -#u #x #k #Hu elim Hu -[ #n #Hn normalize cases (true_or_false (eqb n O)) #H >H [%2] - normalize >(?: leb n O = false) [|cases n in H; // >eqb_n_n #H destruct (H) ] - normalize lapply Hn cases n in H; normalize [ #Hfalse destruct (Hfalse) ] - #n0 #_ #Hn0 % @le_S_S_to_le // -| #x0 %2 -| #v1 #v2 #Hv1 #Hv2 #IH1 #IH2 %3 // -| #x0 #s #v #Hv #IH normalize @daemon -] -qed. -*) - -definition vopen_TM ≝ λu:CTX.λx.mk_TM (u⌈x⌉) ?. -@toc_to_check_tm_true @vopen_tm_down @check_tm_true_to_toc @ctx_of_CTX qed. - -axiom vclose_tm_up : ∀u,x,k.tm_or_ctx k u → tm_or_ctx (S k) (νx.u). - -definition vclose_TM ≝ λu:TM.λx.mk_CTX (νx.u) ?. -@toc_to_check_tm_true @vclose_tm_up @check_tm_true_to_toc @tm_of_TM qed. - -interpretation "ln wf term variable open" 'open u x = (vopen_TM u x). -interpretation "ln wf term variable close" 'nu x u = (vclose_TM u x). - -theorem tm_alpha : ∀x,y,s,u.x ∉ FV u → y ∉ FV u → lam x s (u⌈x⌉) = lam y s (u⌈y⌉). -#x #y #s #u #Hx #Hy whd in ⊢ (??%%); @eq_f >nominal_eta // >nominal_eta // -qed. - -lemma TM_to_tm2 : ∀u:TM.tm2 u. -#u @tm_to_tm2 @check_tm_true_to_toc @tm_of_TM qed. - -theorem TM_ind_plus_weak : - ∀P:pretm → Type[0]. - (∀x:𝔸.P (PAR x)) → - (∀v1,v2:TM.P v1 → P v2 → P (APP v1 v2)) → - ∀C:list 𝔸. - (∀x,s.∀v:CTX.x ∉ FV v → x ∉ C → - (∀y.y ∉ FV v → P (v⌈y⌉)) → P (LAM x s (v⌈x⌉))) → - ∀u:TM.P u. -#P #Hpar #Happ #C #Hlam #u elim (TM_to_tm2 u) // -[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 @(Happ (mk_TM …) (mk_TM …) IH1 IH2) - @toc_to_check_tm_true @tm2_to_tm // -| #x #s #v #Hx #pv #IH - lapply (p_fresh … (C@FV v)) letin x0 ≝ (N_fresh … (C@FV v)) #Hx0 - >(?:lam x s (v⌈x⌉) = lam x0 s (v⌈x0⌉)) - [|@tm_alpha // @(not_to_not … Hx0) @in_list_to_in_list_append_r ] - @(Hlam x0 s (mk_CTX v ?) ??) - [ <(nominal_eta … Hx) @toc_to_check_tm_true @vclose_tm_up @tm2_to_tm @pv // - | @(not_to_not … Hx0) @in_list_to_in_list_append_r - | @(not_to_not … Hx0) @in_list_to_in_list_append_l - | @IH ] -] -qed. - -lemma eq_mk_TM : ∀u,v.u = v → ∀pu,pv.mk_TM u pu = mk_TM v pv. -#u #v #Heq >Heq #pu #pv % -qed. - -lemma eq_P : ∀T:Type[0].∀t1,t2:T.t1 = t2 → ∀P:T → Type[0].P t1 → P t2. // qed. - -theorem TM_ind_plus : - ∀P:TM → Type[0]. - (∀x:𝔸.P (PAR x)) → - (∀v1,v2:TM.P v1 → P v2 → P (APP v1 v2)) → - ∀C:list 𝔸. - (∀x,s.∀v:CTX.x ∉ FV v → x ∉ C → - (∀y.y ∉ FV v → P (v⌈y⌉)) → P (LAM x s (v⌈x⌉))) → - ∀u:TM.P u. -#P #Hpar #Happ #C #Hlam * #u #pu ->(?:mk_TM u pu = - mk_TM u (toc_to_check_tm_true … (tm2_to_tm … (tm_to_tm2 … (check_tm_true_to_toc … pu))))) [|%] -elim (tm_to_tm2 u ?) // -[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 @(Happ (mk_TM …) (mk_TM …) IH1 IH2) -| #x #s #v #Hx #pv #IH - lapply (p_fresh … (C@FV v)) letin x0 ≝ (N_fresh … (C@FV v)) #Hx0 - lapply (Hlam x0 s (mk_CTX v ?) ???) - [2: @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_l @Hx0 - |4: @toc_to_check_tm_true <(nominal_eta x v) // @vclose_tm_up @tm2_to_tm @pv // - | #y #Hy whd in match (vopen_TM ??); - >(?:mk_TM (v⌈y⌉) ? = mk_TM (v⌈y⌉) (toc_to_check_tm_true (v⌈y⌉) O (tm2_to_tm (v⌈y⌉) (pv y Hy)))) - [@IH|%] - | @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_r @Hx0 - | @eq_P @eq_mk_TM whd in match (vopen_TM ??); @tm_alpha // @(not_to_not … Hx0) @in_list_to_in_list_append_r ] -] -qed. - -notation -"hvbox('nominal' u 'return' out 'with' - [ 'xpar' ident x ⇒ f1 - | 'xapp' ident v1 ident v2 ident recv1 ident recv2 ⇒ f2 - | 'xlam' ❨ident y # C❩ ident s ident w ident py1 ident py2 ident recw ⇒ f3 ])" -with precedence 48 -for @{ TM_ind_plus $out (λ${ident x}:?.$f1) - (λ${ident v1}:?.λ${ident v2}:?.λ${ident recv1}:?.λ${ident recv2}:?.$f2) - $C (λ${ident y}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?.λ${ident recw}:?.$f3) - $u }. - -(* include "basics/jmeq.ma".*) - -definition subst ≝ (λu:TM.λx,v. - nominal u return (λ_.TM) with - [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ PAR x0 ] (* u instead of PAR x0 does not work, u stays the same at every rec call! *) - | xapp v1 v2 recv1 recv2 ⇒ APP recv1 recv2 - | xlam ❨y # x::FV v❩ s w py1 py2 recw ⇒ LAM y s (recw y py1) ]). - -lemma subst_def : ∀u,x,v.subst u x v = - nominal u return (λ_.TM) with - [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ PAR x0 ] - | xapp v1 v2 recv1 recv2 ⇒ APP recv1 recv2 - | xlam ❨y # x::FV v❩ s w py1 py2 recw ⇒ LAM y s (recw y py1) ]. // -qed. - -axiom TM_ind_plus_LAM : - ∀x,s,u,out,f1,f2,C,f3,Hx1,Hx2. - TM_ind_plus out f1 f2 C f3 (LAM x s (u⌈x⌉)) = - f3 x s u Hx1 Hx2 (λy,Hy.TM_ind_plus ? f1 f2 C f3 ?). - -axiom TM_ind_plus_APP : - ∀u1,u2,out,f1,f2,C,f3. - TM_ind_plus out f1 f2 C f3 (APP u1 u2) = - f2 u1 u2 (TM_ind_plus out f1 f2 C f3 ?) (TM_ind_plus out f1 f2 C f3 ?). - -lemma eq_mk_CTX : ∀u,v.u = v → ∀pu,pv.mk_CTX u pu = mk_CTX v pv. -#u #v #Heq >Heq #pu #pv % -qed. - -lemma vclose_vopen_TM : ∀x.∀u:TM.((νx.u)⌈x⌉) = u. -#x * #u #pu @eq_mk_TM @vclose_vopen qed. - -lemma nominal_eta_CTX : ∀x.∀u:CTX.x ∉ FV u → (νx.(u⌈x⌉)) = u. -#x * #u #pu #Hx @eq_mk_CTX @nominal_eta // qed. - -theorem TM_alpha : ∀x,y,s.∀u:CTX.x ∉ FV u → y ∉ FV u → LAM x s (u⌈x⌉) = LAM y s (u⌈y⌉). -#x #y #s #u #Hx #Hy @eq_mk_TM @tm_alpha // qed. - -axiom in_vopen_CTX : ∀x,y.∀v:CTX.x ∈ FV (v⌈y⌉) → x = y ∨ x ∈ FV v. - -theorem subst_fresh : ∀u,v:TM.∀x.x ∉ FV u → subst u x v = u. -#u #v #x @(TM_ind_plus … (x::FV v) … u) -[ #x0 normalize in ⊢ (%→?); #Hx normalize in ⊢ (??%?); - >(\bf ?) [| @(not_to_not … Hx) #Heq >Heq % ] % -| #u1 #u2 #IH1 #IH2 normalize in ⊢ (%→?); #Hx - >subst_def >TM_ind_plus_APP @eq_mk_TM @eq_f2 @eq_f - [ subst_def >TM_ind_plus_LAM [|@HC|@Hx0] - @eq_f nominal_eta_CTX // - cases (in_vopen_CTX … Hx) // #Heq >Heq in HC; * #HC @False_ind @HC % -] -qed. - -example subst_LAM_same : ∀x,s,u,v. subst (LAM x s u) x v = LAM x s u. -#x #s #u #v >subst_def <(vclose_vopen_TM x u) -lapply (p_fresh … (FV (νx.u)@x::FV v)) letin x0 ≝ (N_fresh … (FV (νx.u)@x::FV v)) #Hx0 ->(TM_alpha x x0) -[| @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_l @Hx0 | @fresh_vclose_tm ] ->TM_ind_plus_LAM [| @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_r @Hx0 | @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_l @Hx0 ] -@eq_f change with (subst ((νx.u)⌈x0⌉) x v) in ⊢ (??%?); @subst_fresh -@(not_to_not … Hx0) #Hx0' cases (in_vopen_CTX … Hx0') -[ #Heq >Heq @in_list_to_in_list_append_r % -| #Hfalse @False_ind cases (fresh_vclose_tm u x) #H @H @Hfalse ] -qed. - -(* -notation > "Λ ident x. ident T [ident x] ↦ P" - with precedence 48 for @{'foo (λ${ident x}.λ${ident T}.$P)}. - -notation < "Λ ident x. ident T [ident x] ↦ P" - with precedence 48 for @{'foo (λ${ident x}:$Q.λ${ident T}:$R.$P)}. -*) - -(* -notation -"hvbox('nominal' u 'with' - [ 'xpar' ident x ⇒ f1 - | 'xapp' ident v1 ident v2 ⇒ f2 - | 'xlam' ident x # C s w ⇒ f3 ])" -with precedence 48 -for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1) - (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) - $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. -*) - -(* -notation -"hvbox('nominal' u 'with' - [ 'xpar' ident x ^ f1 - | 'xapp' ident v1 ident v2 ^ f2 ])" -(* | 'xlam' ident x # C s w ^ f3 ]) *) -with precedence 48 -for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1) - (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) - $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. -*) -notation -"hvbox('nominal' u 'with' - [ 'xpar' ident x ^ f1 - | 'xapp' ident v1 ident v2 ^ f2 ])" -with precedence 48 -for @{ tm2_ind_plus ? (λ${ident x}:?.$f1) - (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) - $C (λ${ident x}:?.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. - -axiom in_Env : 𝔸 × tp → Env → Prop. -notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}. -interpretation "Env membership" 'lefttriangle x l = (in_Env x l). - - - -inductive judg : list tp → tm → tp → Prop ≝ -| t_var : ∀g,n,t.Nth ? n g = Some ? t → judg g (var n) t -| t_app : ∀g,m,n,t,u.judg g m (arr t u) → judg g n t → judg g (app m n) u -| t_abs : ∀g,t,m,u.judg (t::g) m u → judg g (abs t m) (arr t u). - -definition Env := list (𝔸 × tp). - -axiom vclose_env : Env → list tp. -axiom vclose_tm : Env → tm → tm. -axiom Lam : 𝔸 → tp → tm → tm. -definition Judg ≝ λG,M,T.judg (vclose_env G) (vclose_tm G M) T. -definition dom ≝ λG:Env.map ?? (fst ??) G. - -definition sctx ≝ 𝔸 × tm. -axiom swap_tm : 𝔸 → 𝔸 → tm → tm. -definition sctx_app : sctx → 𝔸 → tm ≝ λM0,Y.let 〈X,M〉 ≝ M0 in swap_tm X Y M. - -axiom in_list : ∀A:Type[0].A → list A → Prop. -interpretation "list membership" 'mem x l = (in_list ? x l). -interpretation "list non-membership" 'notmem x l = (Not (in_list ? x l)). - -axiom in_Env : 𝔸 × tp → Env → Prop. -notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}. -interpretation "Env membership" 'lefttriangle x l = (in_Env x l). - -(* axiom Lookup : 𝔸 → Env → option tp. *) - -(* forma alto livello del judgment - t_abs* : ∀G,T,X,M,U. - (∀Y ∉ supp(M).Judg (〈Y,T〉::G) (M[Y]) U) → - Judg G (Lam X T (M[X])) (arr T U) *) - -(* prima dimostrare, poi perfezionare gli assiomi, poi dimostrarli *) - -axiom Judg_ind : ∀P:Env → tm → tp → Prop. - (∀X,G,T.〈X,T〉 ◃ G → P G (par X) T) → - (∀G,M,N,T,U. - Judg G M (arr T U) → Judg G N T → - P G M (arr T U) → P G N T → P G (app M N) U) → - (∀G,T1,T2,X,M1. - (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → Judg (〈Y,T1〉::G) (sctx_app M1 Y) T2) → - (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → P (〈Y,T1〉::G) (sctx_app M1 Y) T2) → - P G (Lam X T1 (sctx_app M1 X)) (arr T1 T2)) → - ∀G,M,T.Judg G M T → P G M T. - -axiom t_par : ∀X,G,T.〈X,T〉 ◃ G → Judg G (par X) T. -axiom t_app2 : ∀G,M,N,T,U.Judg G M (arr T U) → Judg G N T → Judg G (app M N) U. -axiom t_Lam : ∀G,X,M,T,U.Judg (〈X,T〉::G) M U → Judg G (Lam X T M) (arr T U). - -definition subenv ≝ λG1,G2.∀x.x ◃ G1 → x ◃ G2. -interpretation "subenv" 'subseteq G1 G2 = (subenv G1 G2). - -axiom daemon : ∀P:Prop.P. - -theorem weakening : ∀G1,G2,M,T.G1 ⊆ G2 → Judg G1 M T → Judg G2 M T. -#G1 #G2 #M #T #Hsub #HJ lapply Hsub lapply G2 -G2 change with (∀G2.?) -@(Judg_ind … HJ) -[ #X #G #T0 #Hin #G2 #Hsub @t_par @Hsub // -| #G #M0 #N #T0 #U #HM0 #HN #IH1 #IH2 #G2 #Hsub @t_app2 - [| @IH1 // | @IH2 // ] -| #G #T1 #T2 #X #M1 #HM1 #IH #G2 #Hsub @t_Lam @IH - [ (* trivial property of Lam *) @daemon - | (* trivial property of subenv *) @daemon ] -] -qed. - -(* Serve un tipo Tm per i termini localmente chiusi e i suoi principi di induzione e - ricorsione *) \ No newline at end of file diff --git a/matita/matita/lib/binding/ln_concrete.ma b/matita/matita/lib/binding/ln_concrete.ma deleted file mode 100644 index bdafdb1a0..000000000 --- a/matita/matita/lib/binding/ln_concrete.ma +++ /dev/null @@ -1,683 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basics/lists/list.ma". -include "basics/deqsets.ma". -include "binding/names.ma". -include "binding/fp.ma". - -definition alpha : Nset ≝ X. check alpha -notation "𝔸" non associative with precedence 90 for @{'alphabet}. -interpretation "set of names" 'alphabet = alpha. - -inductive tp : Type[0] ≝ -| top : tp -| arr : tp → tp → tp. -inductive pretm : Type[0] ≝ -| var : nat → pretm -| par : 𝔸 → pretm -| abs : tp → pretm → pretm -| app : pretm → pretm → pretm. - -let rec Nth T n (l:list T) on n ≝ - match l with - [ nil ⇒ None ? - | cons hd tl ⇒ match n with - [ O ⇒ Some ? hd - | S n0 ⇒ Nth T n0 tl ] ]. - -let rec vclose_tm_aux u x k ≝ match u with - [ var n ⇒ if (leb k n) then var (S n) else u - | par x0 ⇒ if (x0 == x) then (var k) else u - | app v1 v2 ⇒ app (vclose_tm_aux v1 x k) (vclose_tm_aux v2 x k) - | abs s v ⇒ abs s (vclose_tm_aux v x (S k)) ]. -definition vclose_tm ≝ λu,x.vclose_tm_aux u x O. - -definition vopen_var ≝ λn,x,k.match eqb n k with - [ true ⇒ par x - | false ⇒ match leb n k with - [ true ⇒ var n - | false ⇒ var (pred n) ] ]. - -let rec vopen_tm_aux u x k ≝ match u with - [ var n ⇒ vopen_var n x k - | par x0 ⇒ u - | app v1 v2 ⇒ app (vopen_tm_aux v1 x k) (vopen_tm_aux v2 x k) - | abs s v ⇒ abs s (vopen_tm_aux v x (S k)) ]. -definition vopen_tm ≝ λu,x.vopen_tm_aux u x O. - -let rec FV u ≝ match u with - [ par x ⇒ [x] - | app v1 v2 ⇒ FV v1@FV v2 - | abs s v ⇒ FV v - | _ ⇒ [ ] ]. - -definition lam ≝ λx,s,u.abs s (vclose_tm u x). - -let rec Pi_map_tm p u on u ≝ match u with -[ par x ⇒ par (p x) -| var _ ⇒ u -| app v1 v2 ⇒ app (Pi_map_tm p v1) (Pi_map_tm p v2) -| abs s v ⇒ abs s (Pi_map_tm p v) ]. - -interpretation "permutation of tm" 'middot p x = (Pi_map_tm p x). - -notation "hvbox(u⌈x⌉)" - with precedence 45 - for @{ 'open $u $x }. - -(* -notation "hvbox(u⌈x⌉)" - with precedence 45 - for @{ 'open $u $x }. -notation "❴ u ❵ x" non associative with precedence 90 for @{ 'open $u $x }. -*) -interpretation "ln term variable open" 'open u x = (vopen_tm u x). -notation < "hvbox(ν x break . u)" - with precedence 20 -for @{'nu $x $u }. -notation > "ν list1 x sep , . term 19 u" with precedence 20 - for ${ fold right @{$u} rec acc @{'nu $x $acc)} }. -interpretation "ln term variable close" 'nu x u = (vclose_tm u x). - -let rec tm_height u ≝ match u with -[ app v1 v2 ⇒ S (max (tm_height v1) (tm_height v2)) -| abs s v ⇒ S (tm_height v) -| _ ⇒ O ]. - -theorem le_n_O_rect_Type0 : ∀n:nat. n ≤ O → ∀P: nat →Type[0]. P O → P n. -#n (cases n) // #a #abs cases (?:False) /2/ qed. - -theorem nat_rect_Type0_1 : ∀n:nat.∀P:nat → Type[0]. -(∀m.(∀p. p < m → P p) → P m) → P n. -#n #P #H -cut (∀q:nat. q ≤ n → P q) /2/ -(elim n) - [#q #HleO (* applica male *) - @(le_n_O_rect_Type0 ? HleO) - @H #p #ltpO cases (?:False) /2/ (* 3 *) - |#p #Hind #q #HleS - @H #a #lta @Hind @le_S_S_to_le /2/ - ] -qed. - -lemma leb_false_to_lt : ∀n,m. leb n m = false → m < n. -#n elim n -[ #m normalize #H destruct(H) -| #n0 #IH * // #m normalize #H @le_S_S @IH // ] -qed. - -lemma nominal_eta_aux : ∀x,u.x ∉ FV u → ∀k.vclose_tm_aux (vopen_tm_aux u x k) x k = u. -#x #u elim u -[ #n #_ #k normalize cases (decidable_eq_nat n k) #Hnk - [ >Hnk >eqb_n_n whd in ⊢ (??%?); >(\b ?) // - | >(not_eq_to_eqb_false … Hnk) normalize cases (true_or_false (leb n k)) #Hleb - [ >Hleb normalize >(?:leb k n = false) // - @lt_to_leb_false @not_eq_to_le_to_lt /2/ - | >Hleb normalize >(?:leb k (pred n) = true) normalize - [ cases (leb_false_to_lt … Hleb) // - | @le_to_leb_true cases (leb_false_to_lt … Hleb) normalize /2/ ] ] ] -| #y normalize in ⊢ (%→?→?); #Hy whd in ⊢ (?→??%?); >(\bf ?) // @(not_to_not … Hy) // -| #s #v #IH normalize #Hv #k >IH // @Hv -| #v1 #v2 #IH1 #IH2 normalize #Hv1v2 #k - >IH1 [ >IH2 // | @(not_to_not … Hv1v2) @in_list_to_in_list_append_l ] - @(not_to_not … Hv1v2) @in_list_to_in_list_append_r ] -qed. - -corollary nominal_eta : ∀x,u.x ∉ FV u → (νx.u⌈x⌉) = u. -#x #u #Hu @nominal_eta_aux // -qed. - -lemma eq_height_vopen_aux : ∀v,x,k.tm_height (vopen_tm_aux v x k) = tm_height v. -#v #x elim v -[ #n #k normalize cases (eqb n k) // cases (leb n k) // -| #u #k % -| #s #u #IH #k normalize >IH % -| #u1 #u2 #IH1 #IH2 #k normalize >IH1 >IH2 % ] -qed. - -corollary eq_height_vopen : ∀v,x.tm_height (v⌈x⌉) = tm_height v. -#v #x @eq_height_vopen_aux -qed. - -theorem pretm_ind_plus_aux : - ∀P:pretm → Type[0]. - (∀x:𝔸.P (par x)) → - (∀n:ℕ.P (var n)) → - (∀v1,v2. P v1 → P v2 → P (app v1 v2)) → - ∀C:list 𝔸. - (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) → - ∀n,u.tm_height u ≤ n → P u. -#P #Hpar #Hvar #Happ #C #Hlam #n change with ((λn.?) n); @(nat_rect_Type0_1 n ??) -#m cases m -[ #_ * /2/ - [ normalize #s #v #Hfalse cases (?:False) cases (not_le_Sn_O (tm_height v)) /2/ - | #v1 #v2 whd in ⊢ (?%?→?); #Hfalse cases (?:False) cases (not_le_Sn_O (S (max ??))) /2/ ] ] --m #m #IH * /2/ -[ #s #v whd in ⊢ (?%?→?); #Hv - lapply (p_fresh … (C@FV v)) letin y ≝ (N_fresh … (C@FV v)) #Hy - >(?:abs s v = lam y s (v⌈y⌉)) - [| whd in ⊢ (???%); >nominal_eta // @(not_to_not … Hy) @in_list_to_in_list_append_r ] - @Hlam - [ @(not_to_not … Hy) @in_list_to_in_list_append_r - | @(not_to_not … Hy) @in_list_to_in_list_append_l ] - @IH [| @Hv | >eq_height_vopen % ] -| #v1 #v2 whd in ⊢ (?%?→?); #Hv @Happ - [ @IH [| @Hv | // ] | @IH [| @Hv | // ] ] ] -qed. - -corollary pretm_ind_plus : - ∀P:pretm → Type[0]. - (∀x:𝔸.P (par x)) → - (∀n:ℕ.P (var n)) → - (∀v1,v2. P v1 → P v2 → P (app v1 v2)) → - ∀C:list 𝔸. - (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) → - ∀u.P u. -#P #Hpar #Hvar #Happ #C #Hlam #u @pretm_ind_plus_aux /2/ -qed. - -(* maps a permutation to a list of terms *) -definition Pi_map_list : (𝔸 → 𝔸) → list 𝔸 → list 𝔸 ≝ map 𝔸 𝔸 . - -(* interpretation "permutation of name list" 'middot p x = (Pi_map_list p x).*) - -(* -inductive tm : pretm → Prop ≝ -| tm_par : ∀x:𝔸.tm (par x) -| tm_app : ∀u,v.tm u → tm v → tm (app u v) -| tm_lam : ∀x,s,u.tm u → tm (lam x s u). - -inductive ctx_aux : nat → pretm → Prop ≝ -| ctx_var : ∀n,k.n < k → ctx_aux k (var n) -| ctx_par : ∀x,k.ctx_aux k (par x) -| ctx_app : ∀u,v,k.ctx_aux k u → ctx_aux k v → ctx_aux k (app u v) -(* è sostituibile da ctx_lam ? *) -| ctx_abs : ∀s,u.ctx_aux (S k) u → ctx_aux k (abs s u). -*) - -inductive tm_or_ctx (k:nat) : pretm → Type[0] ≝ -| toc_var : ∀n.n < k → tm_or_ctx k (var n) -| toc_par : ∀x.tm_or_ctx k (par x) -| toc_app : ∀u,v.tm_or_ctx k u → tm_or_ctx k v → tm_or_ctx k (app u v) -| toc_lam : ∀x,s,u.tm_or_ctx k u → tm_or_ctx k (lam x s u). - -definition tm ≝ λt.tm_or_ctx O t. -definition ctx ≝ λt.tm_or_ctx 1 t. - -record TM : Type[0] ≝ { - pretm_of_TM :> pretm; - tm_of_TM : tm pretm_of_TM -}. - -record CTX : Type[0] ≝ { - pretm_of_CTX :> pretm; - ctx_of_CTX : ctx pretm_of_CTX -}. - -inductive tm2 : pretm → Type[0] ≝ -| tm_par : ∀x.tm2 (par x) -| tm_app : ∀u,v.tm2 u → tm2 v → tm2 (app u v) -| tm_lam : ∀x,s,u.x ∉ FV u → (∀y.y ∉ FV u → tm2 (u⌈y⌉)) → tm2 (lam x s (u⌈x⌉)). - -(* -inductive tm' : pretm → Prop ≝ -| tm_par : ∀x.tm' (par x) -| tm_app : ∀u,v.tm' u → tm' v → tm' (app u v) -| tm_lam : ∀x,s,u,C.x ∉ FV u → x ∉ C → (∀y.y ∉ FV u → tm' (❴u❵y)) → tm' (lam x s (❴u❵x)). -*) - -axiom swap_inj : ∀N.∀z1,z2,x,y.swap N z1 z2 x = swap N z1 z2 y → x = y. - -lemma pi_vclose_tm : - ∀z1,z2,x,u.swap 𝔸 z1 z2·(νx.u) = (ν swap ? z1 z2 x.swap 𝔸 z1 z2 · u). -#z1 #z2 #x #u -change with (vclose_tm_aux ???) in match (vclose_tm ??); -change with (vclose_tm_aux ???) in ⊢ (???%); lapply O elim u -[3:whd in ⊢ (?→?→(?→ ??%%)→?→??%%); // -|4:whd in ⊢ (?→?→(?→??%%)→(?→??%%)→?→??%%); // -| #n #k whd in ⊢ (??(??%)%); cases (leb k n) normalize % -| #x0 #k cases (true_or_false (x0==z1)) #H1 >H1 whd in ⊢ (??%%); - [ cases (true_or_false (x0==x)) #H2 >H2 whd in ⊢ (??(??%)%); - [ <(\P H2) >H1 whd in ⊢ (??(??%)%); >(\b ?) // >(\b ?) // - | >H2 whd in match (swap ????); >H1 - whd in match (if false then var k else ?); - whd in match (if true then z2 else ?); >(\bf ?) - [ >(\P H1) >swap_left % - | <(swap_inv ? z1 z2 z2) in ⊢ (?(??%?)); % #H3 - lapply (swap_inj … H3) >swap_right #H4

H1 #H destruct (H) ] - ] - | >(?:(swap ? z1 z2 x0 == swap ? z1 z2 x) = (x0 == x)) - [| cases (true_or_false (x0==x)) #H2 >H2 - [ >(\P H2) @(\b ?) % - | @(\bf ?) % #H >(swap_inj … H) in H2; >(\b ?) // #H0 destruct (H0) ] ] - cases (true_or_false (x0==x)) #H2 >H2 whd in ⊢ (??(??%)%); - [ <(\P H2) >H1 >(\b (refl ??)) % - | >H1 >H2 % ] - ] - ] -qed. - -lemma pi_vopen_tm : - ∀z1,z2,x,u.swap 𝔸 z1 z2·(u⌈x⌉) = (swap 𝔸 z1 z2 · u⌈swap 𝔸 z1 z2 x⌉). -#z1 #z2 #x #u -change with (vopen_tm_aux ???) in match (vopen_tm ??); -change with (vopen_tm_aux ???) in ⊢ (???%); lapply O elim u // -[2: #s #v whd in ⊢ ((?→??%%)→?→??%%); // -|3: #v1 #v2 whd in ⊢ ((?→??%%)→(?→??%%)→?→??%%); /2/ ] -#n #k whd in ⊢ (??(??%)%); cases (true_or_false (eqb n k)) #H1 >H1 // -cases (true_or_false (leb n k)) #H2 >H2 normalize // -qed. - -lemma pi_lam : - ∀z1,z2,x,s,u.swap 𝔸 z1 z2 · lam x s u = lam (swap 𝔸 z1 z2 x) s (swap 𝔸 z1 z2 · u). -#z1 #z2 #x #s #u whd in ⊢ (???%); <(pi_vclose_tm …) % -qed. - -lemma eqv_FV : ∀z1,z2,u.FV (swap 𝔸 z1 z2 · u) = Pi_map_list (swap 𝔸 z1 z2) (FV u). -#z1 #z2 #u elim u // -[ #s #v #H @H -| #v1 #v2 whd in ⊢ (??%%→??%%→??%%); #H1 #H2 >H1 >H2 - whd in ⊢ (???(????%)); /2/ ] -qed. - -lemma swap_inv_tm : ∀z1,z2,u.swap 𝔸 z1 z2 · (swap 𝔸 z1 z2 · u) = u. -#z1 #z2 #u elim u -[1: #n % -|3: #s #v whd in ⊢ (?→??%%); // -|4: #v1 #v2 #Hv1 #Hv2 whd in ⊢ (??%%); // ] -#x whd in ⊢ (??%?); >swap_inv % -qed. - -lemma eqv_in_list : ∀x,l,z1,z2.x ∈ l → swap 𝔸 z1 z2 x ∈ Pi_map_list (swap 𝔸 z1 z2) l. -#x #l #z1 #z2 #Hin elim Hin -[ #x0 #l0 % -| #x1 #x2 #l0 #Hin #IH %2 @IH ] -qed. - -lemma eqv_tm2 : ∀u.tm2 u → ∀z1,z2.tm2 ((swap ? z1 z2)·u). -#u #Hu #z1 #z2 letin p ≝ (swap ? z1 z2) elim Hu /2/ -#x #s #v #Hx #Hv #IH >pi_lam >pi_vopen_tm %3 -[ @(not_to_not … Hx) -Hx #Hx - <(swap_inv ? z1 z2 x) <(swap_inv_tm z1 z2 v) >eqv_FV @eqv_in_list // -| #y #Hy <(swap_inv ? z1 z2 y) - eqv_FV @eqv_in_list // -] -qed. - -lemma vclose_vopen_aux : ∀x,u,k.vopen_tm_aux (vclose_tm_aux u x k) x k = u. -#x #u elim u [1,3,4:normalize //] -[ #n #k cases (true_or_false (leb k n)) #H >H whd in ⊢ (??%?); - [ cases (true_or_false (eqb (S n) k)) #H1 >H1 - [ <(eqb_true_to_eq … H1) in H; #H lapply (leb_true_to_le … H) -H #H - cases (le_to_not_lt … H) -H #H cases (H ?) % - | whd in ⊢ (??%?); >lt_to_leb_false // @le_S_S /2/ ] - | cases (true_or_false (eqb n k)) #H1 >H1 normalize - [ >(eqb_true_to_eq … H1) in H; #H lapply (leb_false_to_not_le … H) -H - * #H cases (H ?) % - | >le_to_leb_true // @not_lt_to_le % #H2 >le_to_leb_true in H; - [ #H destruct (H) | /2/ ] - ] - ] -| #x0 #k whd in ⊢ (??(?%??)?); cases (true_or_false (x0==x)) - #H1 >H1 normalize // >(\P H1) >eqb_n_n % ] -qed. - -lemma vclose_vopen : ∀x,u.((νx.u)⌈x⌉) = u. #x #u @vclose_vopen_aux -qed. - -(* -theorem tm_to_tm : ∀t.tm' t → tm t. -#t #H elim H -*) - -lemma in_list_singleton : ∀T.∀t1,t2:T.t1 ∈ [t2] → t1 = t2. -#T #t1 #t2 #H @(in_list_inv_ind ??? H) /2/ -qed. - -lemma fresh_vclose_tm_aux : ∀u,x,k.x ∉ FV (vclose_tm_aux u x k). -#u #x elim u // -[ #n #k normalize cases (leb k n) normalize // -| #x0 #k whd in ⊢ (?(???(?%))); cases (true_or_false (x0==x)) #H >H normalize // - lapply (\Pf H) @not_to_not #Hin >(in_list_singleton ??? Hin) % -| #v1 #v2 #IH1 #IH2 #k normalize % #Hin cases (in_list_append_to_or_in_list ???? Hin) -Hin #Hin - [ cases (IH1 k) -IH1 #IH1 @IH1 @Hin | cases (IH2 k) -IH2 #IH2 @IH2 @Hin ] -qed. - -lemma fresh_vclose_tm : ∀u,x.x ∉ FV (νx.u). // -qed. - -lemma fresh_swap_tm : ∀z1,z2,u.z1 ∉ FV u → z2 ∉ FV u → swap 𝔸 z1 z2 · u = u. -#z1 #z2 #u elim u -[2: normalize in ⊢ (?→%→%→?); #x #Hz1 #Hz2 whd in ⊢ (??%?); >swap_other // - [ @(not_to_not … Hz2) | @(not_to_not … Hz1) ] // -|1: // -| #s #v #IH normalize #Hz1 #Hz2 >IH // [@Hz2|@Hz1] -| #v1 #v2 #IH1 #IH2 normalize #Hz1 #Hz2 - >IH1 [| @(not_to_not … Hz2) @in_list_to_in_list_append_l | @(not_to_not … Hz1) @in_list_to_in_list_append_l ] - >IH2 // [@(not_to_not … Hz2) @in_list_to_in_list_append_r | @(not_to_not … Hz1) @in_list_to_in_list_append_r ] -] -qed. - -theorem tm_to_tm2 : ∀u.tm u → tm2 u. -#t #Ht elim Ht -[ #n #Hn cases (not_le_Sn_O n) #Hfalse cases (Hfalse Hn) -| @tm_par -| #u #v #Hu #Hv @tm_app -| #x #s #u #Hu #IHu <(vclose_vopen x u) @tm_lam - [ @fresh_vclose_tm - | #y #Hy <(fresh_swap_tm x y (νx.u)) /2/ @fresh_vclose_tm ] -] -qed. - -theorem tm2_to_tm : ∀u.tm2 u → tm u. -#u #pu elim pu /2/ #x #s #v #Hx #Hv #IH %4 @IH // -qed. - -definition PAR ≝ λx.mk_TM (par x) ?. // qed. -definition APP ≝ λu,v:TM.mk_TM (app u v) ?./2/ qed. -definition LAM ≝ λx,s.λu:TM.mk_TM (lam x s u) ?./2/ qed. - -axiom vopen_tm_down : ∀u,x,k.tm_or_ctx (S k) u → tm_or_ctx k (u⌈x⌉). -(* needs true_plus_false - -#u #x #k #Hu elim Hu -[ #n #Hn normalize cases (true_or_false (eqb n O)) #H >H [%2] - normalize >(?: leb n O = false) [|cases n in H; // >eqb_n_n #H destruct (H) ] - normalize lapply Hn cases n in H; normalize [ #Hfalse destruct (Hfalse) ] - #n0 #_ #Hn0 % @le_S_S_to_le // -| #x0 %2 -| #v1 #v2 #Hv1 #Hv2 #IH1 #IH2 %3 // -| #x0 #s #v #Hv #IH normalize @daemon -] -qed. -*) - -definition vopen_TM ≝ λu:CTX.λx.mk_TM (u⌈x⌉) (vopen_tm_down …). @ctx_of_CTX qed. - -axiom vclose_tm_up : ∀u,x,k.tm_or_ctx k u → tm_or_ctx (S k) (νx.u). - -definition vclose_TM ≝ λu:TM.λx.mk_CTX (νx.u) (vclose_tm_up …). @tm_of_TM qed. - -interpretation "ln wf term variable open" 'open u x = (vopen_TM u x). -interpretation "ln wf term variable close" 'nu x u = (vclose_TM u x). - -theorem tm_alpha : ∀x,y,s,u.x ∉ FV u → y ∉ FV u → lam x s (u⌈x⌉) = lam y s (u⌈y⌉). -#x #y #s #u #Hx #Hy whd in ⊢ (??%%); @eq_f >nominal_eta // >nominal_eta // -qed. - -theorem TM_ind_plus : -(* non si può dare il principio in modo dipendente (almeno utilizzando tm2) - la "prova" purtroppo è in Type e non si può garantire che sia esattamente - quella che ci aspetteremmo - *) - ∀P:pretm → Type[0]. - (∀x:𝔸.P (PAR x)) → - (∀v1,v2:TM.P v1 → P v2 → P (APP v1 v2)) → - ∀C:list 𝔸. - (∀x,s.∀v:CTX.x ∉ FV v → x ∉ C → - (∀y.y ∉ FV v → P (v⌈y⌉)) → P (LAM x s (v⌈x⌉))) → - ∀u:TM.P u. -#P #Hpar #Happ #C #Hlam * #u #pu elim (tm_to_tm2 u pu) // -[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 @(Happ (mk_TM …) (mk_TM …)) /2/ -| #x #s #v #Hx #pv #IH - lapply (p_fresh … (C@FV v)) letin x0 ≝ (N_fresh … (C@FV v)) #Hx0 - >(?:lam x s (v⌈x⌉) = lam x0 s (v⌈x0⌉)) - [|@tm_alpha // @(not_to_not … Hx0) @in_list_to_in_list_append_r ] - @(Hlam x0 s (mk_CTX v ?) ??) - [ <(nominal_eta … Hx) @vclose_tm_up @tm2_to_tm @pv // - | @(not_to_not … Hx0) @in_list_to_in_list_append_r - | @(not_to_not … Hx0) @in_list_to_in_list_append_l - | @IH ] -] -qed. - -notation -"hvbox('nominal' u 'return' out 'with' - [ 'xpar' ident x ⇒ f1 - | 'xapp' ident v1 ident v2 ident recv1 ident recv2 ⇒ f2 - | 'xlam' ❨ident y # C❩ ident s ident w ident py1 ident py2 ident recw ⇒ f3 ])" -with precedence 48 -for @{ TM_ind_plus $out (λ${ident x}:?.$f1) - (λ${ident v1}:?.λ${ident v2}:?.λ${ident recv1}:?.λ${ident recv2}:?.$f2) - $C (λ${ident y}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?.λ${ident recw}:?.$f3) - $u }. - -(* include "basics/jmeq.ma".*) - -definition subst ≝ (λu:TM.λx,v. - nominal u return (λ_.TM) with - [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ u ] - | xapp v1 v2 recv1 recv2 ⇒ APP recv1 recv2 - | xlam ❨y # x::FV v❩ s w py1 py2 recw ⇒ LAM y s (recw y py1) ]). - -lemma fasfd : ∀s,v. pretm_of_TM (subst (LAM O s (PAR 1)) O v) = pretm_of_TM (LAM O s (PAR 1)). -#s #v normalize in ⊢ (??%?); - - -theorem tm2_ind_plus : -(* non si può dare il principio in modo dipendente (almeno utilizzando tm2) *) - ∀P:pretm → Type[0]. - (∀x:𝔸.P (par x)) → - (∀v1,v2.tm2 v1 → tm2 v2 → P v1 → P v2 → P (app v1 v2)) → - ∀C:list 𝔸. - (∀x,s,v.x ∉ FV v → x ∉ C → (∀y.y ∉ FV v → tm2 (v⌈y⌉)) → - (∀y.y ∉ FV v → P (v⌈y⌉)) → P (lam x s (v⌈x⌉))) → - ∀u.tm2 u → P u. -#P #Hpar #Happ #C #Hlam #u #pu elim pu /2/ -#x #s #v #px #pv #IH -lapply (p_fresh … (C@FV v)) letin y ≝ (N_fresh … (C@FV v)) #Hy ->(?:lam x s (v⌈x⌉) = lam y s (v⌈y⌉)) [| @tm_alpha // @(not_to_not … Hy) @in_list_to_in_list_append_r ] -@Hlam /2/ lapply Hy -Hy @not_to_not #Hy -[ @in_list_to_in_list_append_r @Hy | @in_list_to_in_list_append_l @Hy ] -qed. - -definition check_tm ≝ - λu.pretm_ind_plus ? (λ_.true) (λ_.false) - (λv1,v2,r1,r2.r1 ∧ r2) [ ] (λx,s,v,pv1,pv2,rv.rv) u. - -(* -lemma check_tm_complete : ∀u.tm u → check_tm u = true. -#u #pu @(tm2_ind_plus … [ ] … (tm_to_tm2 ? pu)) // -[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 -| #x #s #v #Hx1 #Hx2 #Hv #IH -*) - -notation -"hvbox('nominal' u 'return' out 'with' - [ 'xpar' ident x ⇒ f1 - | 'xapp' ident v1 ident v2 ident pv1 ident pv2 ident recv1 ident recv2 ⇒ f2 - | 'xlam' ❨ident y # C❩ ident s ident w ident py1 ident py2 ident pw ident recw ⇒ f3 ])" -with precedence 48 -for @{ tm2_ind_plus $out (λ${ident x}:?.$f1) - (λ${ident v1}:?.λ${ident v2}:?.λ${ident pv1}:?.λ${ident pv2}:?.λ${ident recv1}:?.λ${ident recv2}:?.$f2) - $C (λ${ident y}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?.λ${ident pw}:?.λ${ident recw}:?.$f3) - ? (tm_to_tm2 ? $u) }. -(* notation -"hvbox('nominal' u 'with' - [ 'xlam' ident x # C ident s ident w ⇒ f3 ])" -with precedence 48 -for @{ tm2_ind_plus ??? - $C (λ${ident x}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?. - λ${ident pw}:?.λ${ident recw}:?.$f3) $u (tm_to_tm2 ??) }. -*) - - -definition subst ≝ (λu.λpu:tm u.λx,v. - nominal pu return (λ_.pretm) with - [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ u ] - | xapp v1 v2 pv1 pv2 recv1 recv2 ⇒ app recv1 recv2 - | xlam ❨y # x::FV v❩ s w py1 py2 pw recw ⇒ lam y s (recw y py1) ]). - -lemma fasfd : ∀x,s,u,p1,v. subst (lam x s u) p1 x v = lam x s u. -#x #s #u #p1 #v - - -definition subst ≝ λu.λpu:tm u.λx,y. - tm2_ind_plus ? - (* par x0 *) (λx0.match x == x0 with [ true ⇒ v | false ⇒ u ]) - (* app v1 v2 *) (λv1,v2,pv1,pv2,recv1,recv2.app recv1 recv2) - (* lam y#(x::FV v) s w *) (x::FV v) (λy,s,w,py1,py2,pw,recw.lam y s (recw y py1)) - u (tm_to_tm2 … pu). -check subst -definition subst ≝ λu.λpu:tm u.λx,v. - nominal u with - [ xlam y # (x::FV v) s w ^ ? ]. - -(* -notation > "Λ ident x. ident T [ident x] ↦ P" - with precedence 48 for @{'foo (λ${ident x}.λ${ident T}.$P)}. - -notation < "Λ ident x. ident T [ident x] ↦ P" - with precedence 48 for @{'foo (λ${ident x}:$Q.λ${ident T}:$R.$P)}. -*) - -(* -notation -"hvbox('nominal' u 'with' - [ 'xpar' ident x ⇒ f1 - | 'xapp' ident v1 ident v2 ⇒ f2 - | 'xlam' ident x # C s w ⇒ f3 ])" -with precedence 48 -for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1) - (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) - $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. -*) - -(* -notation -"hvbox('nominal' u 'with' - [ 'xpar' ident x ^ f1 - | 'xapp' ident v1 ident v2 ^ f2 ])" -(* | 'xlam' ident x # C s w ^ f3 ]) *) -with precedence 48 -for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1) - (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) - $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. -*) -notation -"hvbox('nominal' u 'with' - [ 'xpar' ident x ^ f1 - | 'xapp' ident v1 ident v2 ^ f2 ])" -with precedence 48 -for @{ tm2_ind_plus ? (λ${ident x}:?.$f1) - (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2) - $C (λ${ident x}:?.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }. - - -definition subst ≝ λu.λpu:tm u.λx,v. - nominal u with - [ xpar x0 ^ match x == x0 with [ true ⇒ v | false ⇒ u ] - | xapp v1 v2 ^ ? ]. - | xlam y # (x::FV v) s w ^ ? ]. - - - (* par x0 *) (λx0.match x == x0 with [ true ⇒ v | false ⇒ u ]) - (* app v1 v2 *) (λv1,v2,pv1,pv2,recv1,recv2.app recv1 recv2) - (* lam y#(x::FV v) s w *) (x::FV v) (λy,s,w,py1,py2,pw,recw.lam y s (recw y py1)) - u (tm_to_tm2 … pu). - - -*) -definition subst ≝ λu.λpu:tm u.λx,v. - tm2_ind_plus ? - (* par x0 *) (λx0.match x == x0 with [ true ⇒ v | false ⇒ u ]) - (* app v1 v2 *) (λv1,v2,pv1,pv2,recv1,recv2.app recv1 recv2) - (* lam y#(x::FV v) s w *) (x::FV v) (λy,s,w,py1,py2,pw,recw.lam y s (recw y py1)) - u (tm_to_tm2 … pu). - -check subst - - -axiom in_Env : 𝔸 × tp → Env → Prop. -notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}. -interpretation "Env membership" 'lefttriangle x l = (in_Env x l). - - - -inductive judg : list tp → tm → tp → Prop ≝ -| t_var : ∀g,n,t.Nth ? n g = Some ? t → judg g (var n) t -| t_app : ∀g,m,n,t,u.judg g m (arr t u) → judg g n t → judg g (app m n) u -| t_abs : ∀g,t,m,u.judg (t::g) m u → judg g (abs t m) (arr t u). - -definition Env := list (𝔸 × tp). - -axiom vclose_env : Env → list tp. -axiom vclose_tm : Env → tm → tm. -axiom Lam : 𝔸 → tp → tm → tm. -definition Judg ≝ λG,M,T.judg (vclose_env G) (vclose_tm G M) T. -definition dom ≝ λG:Env.map ?? (fst ??) G. - -definition sctx ≝ 𝔸 × tm. -axiom swap_tm : 𝔸 → 𝔸 → tm → tm. -definition sctx_app : sctx → 𝔸 → tm ≝ λM0,Y.let 〈X,M〉 ≝ M0 in swap_tm X Y M. - -axiom in_list : ∀A:Type[0].A → list A → Prop. -interpretation "list membership" 'mem x l = (in_list ? x l). -interpretation "list non-membership" 'notmem x l = (Not (in_list ? x l)). - -axiom in_Env : 𝔸 × tp → Env → Prop. -notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}. -interpretation "Env membership" 'lefttriangle x l = (in_Env x l). - -let rec FV M ≝ match M with - [ par X ⇒ [X] - | app M1 M2 ⇒ FV M1@FV M2 - | abs T M0 ⇒ FV M0 - | _ ⇒ [ ] ]. - -(* axiom Lookup : 𝔸 → Env → option tp. *) - -(* forma alto livello del judgment - t_abs* : ∀G,T,X,M,U. - (∀Y ∉ supp(M).Judg (〈Y,T〉::G) (M[Y]) U) → - Judg G (Lam X T (M[X])) (arr T U) *) - -(* prima dimostrare, poi perfezionare gli assiomi, poi dimostrarli *) - -axiom Judg_ind : ∀P:Env → tm → tp → Prop. - (∀X,G,T.〈X,T〉 ◃ G → P G (par X) T) → - (∀G,M,N,T,U. - Judg G M (arr T U) → Judg G N T → - P G M (arr T U) → P G N T → P G (app M N) U) → - (∀G,T1,T2,X,M1. - (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → Judg (〈Y,T1〉::G) (sctx_app M1 Y) T2) → - (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → P (〈Y,T1〉::G) (sctx_app M1 Y) T2) → - P G (Lam X T1 (sctx_app M1 X)) (arr T1 T2)) → - ∀G,M,T.Judg G M T → P G M T. - -axiom t_par : ∀X,G,T.〈X,T〉 ◃ G → Judg G (par X) T. -axiom t_app2 : ∀G,M,N,T,U.Judg G M (arr T U) → Judg G N T → Judg G (app M N) U. -axiom t_Lam : ∀G,X,M,T,U.Judg (〈X,T〉::G) M U → Judg G (Lam X T M) (arr T U). - -definition subenv ≝ λG1,G2.∀x.x ◃ G1 → x ◃ G2. -interpretation "subenv" 'subseteq G1 G2 = (subenv G1 G2). - -axiom daemon : ∀P:Prop.P. - -theorem weakening : ∀G1,G2,M,T.G1 ⊆ G2 → Judg G1 M T → Judg G2 M T. -#G1 #G2 #M #T #Hsub #HJ lapply Hsub lapply G2 -G2 change with (∀G2.?) -@(Judg_ind … HJ) -[ #X #G #T0 #Hin #G2 #Hsub @t_par @Hsub // -| #G #M0 #N #T0 #U #HM0 #HN #IH1 #IH2 #G2 #Hsub @t_app2 - [| @IH1 // | @IH2 // ] -| #G #T1 #T2 #X #M1 #HM1 #IH #G2 #Hsub @t_Lam @IH - [ (* trivial property of Lam *) @daemon - | (* trivial property of subenv *) @daemon ] -] -qed. - -(* Serve un tipo Tm per i termini localmente chiusi e i suoi principi di induzione e - ricorsione *) \ No newline at end of file diff --git a/matita/matita/lib/binding/names.ma b/matita/matita/lib/binding/names.ma deleted file mode 100644 index bc5f07d43..000000000 --- a/matita/matita/lib/binding/names.ma +++ /dev/null @@ -1,78 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basics/logic.ma". -include "basics/lists/in.ma". -include "basics/types.ma". - -(*interpretation "list membership" 'mem x l = (in_list ? x l).*) - -record Nset : Type[1] ≝ -{ - (* carrier is specified as a coercion: when an object X of type Nset is - given, but something of type Type is expected, Matita will insert a - hidden coercion: the user sees "X", but really means "carrier X" *) - carrier :> DeqSet; - N_fresh : list carrier → carrier; - p_fresh : ∀l.N_fresh l ∉ l -}. - -definition maxlist ≝ - λl.foldr ?? (λx,acc.max x acc) 0 l. - -definition natfresh ≝ λl.S (maxlist l). - -lemma le_max_1 : ∀x,y.x ≤ max x y. /2/ -qed. - -lemma le_max_2 : ∀x,y.y ≤ max x y. /2/ -qed. - -lemma le_maxlist : ∀l,x.x ∈ l → x ≤ maxlist l. -#l elim l -[#x #Hx @False_ind cases (not_in_list_nil ? x) #H1 /2/ -|#y #tl #IH #x #H1 change with (max ??) in ⊢ (??%); - cases (in_list_cons_case ???? H1);#H2; - [ >H2 @le_max_1 - | whd in ⊢ (??%); lapply (refl ? (leb y (maxlist tl))); - cases (leb y (maxlist tl)) in ⊢ (???% → %);#H3 - [ @IH // - | lapply (IH ? H2) #H4 - lapply (leb_false_to_not_le … H3) #H5 - lapply (not_le_to_lt … H5) #H6 - @(transitive_le … H4) - @(transitive_le … H6) %2 % - ] - ] -] -qed. - -(* prove freshness for nat *) -lemma lt_l_natfresh_l : ∀l,x.x ∈ l → x < natfresh l. -#l #x #H1 @le_S_S /2/ -qed. - -(*naxiom p_Xfresh : ∀l.∀x:Xcarr.x ∈ l → x ≠ ntm (Xfresh l) ∧ x ≠ ntp (Xfresh l).*) -lemma p_natfresh : ∀l.natfresh l ∉ l. -#l % #H1 lapply (lt_l_natfresh_l … H1) #H2 -cases (lt_to_not_eq … H2) #H3 @H3 % -qed. - -include "basics/finset.ma". - -definition X : Nset ≝ mk_Nset DeqNat …. -[ @natfresh -| @p_natfresh -] -qed. \ No newline at end of file diff --git a/matita/matita/lib/finite_lambda/confluence.ma b/matita/matita/lib/finite_lambda/confluence.ma deleted file mode 100644 index 32b7e3ff0..000000000 --- a/matita/matita/lib/finite_lambda/confluence.ma +++ /dev/null @@ -1,226 +0,0 @@ -(* - ||M|| This file is part of HELM, an Hypertextual, Electronic - ||A|| Library of Mathematics, developed at the Computer Science - ||T|| Department of the University of Bologna, Italy. - ||I|| - ||T|| - ||A|| This file is distributed under the terms of the - \ / GNU General Public License Version 2 - \ / - V_______________________________________________________________ *) - -include "finite_lambda/reduction.ma". - - -axiom canonical_to_T: ∀O,D.∀M:T O D.∀ty.(* type_of M ty → *) - ∃a:FinSet_of_FType O D ty. star ? (red O D) M (to_T O D ty a). - -axiom normal_to_T: ∀O,D,M,ty,a. red O D (to_T O D ty a) M → False. - -axiom red_closed: ∀O,D,M,M1. - is_closed O D 0 M → red O D M M1 → is_closed O D 0 M1. - -lemma critical: ∀O,D,ty,M,N. - ∃M3:T O D - .star (T O D) (red O D) (subst O D M 0 N) M3 - ∧star (T O D) (red O D) - (App O D - (Vec O D ty - (map (FinSet_of_FType O D ty) (T O D) - (λa0:FinSet_of_FType O D ty.subst O D M 0 (to_T O D ty a0)) - (enum (FinSet_of_FType O D ty)))) N) M3. -#O #D #ty #M #N -lapply (canonical_to_T O D N ty) * #a #Ha -%{(subst O D M 0 (to_T O D ty a))} (* CR-term *) -%[@red_star_subst @Ha - |@trans_star [|@(star_red_appr … Ha)] @R_to_star @riota - lapply (enum_complete (FinSet_of_FType O D ty) a) - elim (enum (FinSet_of_FType O D ty)) - [normalize #H1 destruct (H1) - |#hd #tl #Hind #H cases (orb_true_l … H) -H #Hcase - [normalize >Hcase >(\P Hcase) // - |normalize cases (true_or_false (a==hd)) #Hcase1 - [normalize >Hcase1 >(\P Hcase1) // |>Hcase1 @Hind @Hcase] - ] - ] - ] -qed. - -lemma critical2: ∀O,D,ty,a,M,M1,M2,v. - red O D (Vec O D ty v) M → - red O D (App O D (Vec O D ty v) (to_T O D ty a)) M1 → - assoc (FinSet_of_FType O D ty) (T O D) a (enum (FinSet_of_FType O D ty)) v - =Some (T O D) M2 → - ∃M3:T O D - .star (T O D) (red O D) M2 M3 - ∧star (T O D) (red O D) (App O D M (to_T O D ty a)) M3. -#O #D #ty #a #M #M1 #M2 #v #redM #redM1 #Ha lapply (red_vec … redM) -redM -* #N * #N1 * #v1 * #v2 * * #Hred1 #Hv #HM0 >HM0 -HM0 >Hv in Ha; #Ha -cases (same_assoc … a (enum (FinSet_of_FType O D ty)) v1 v2 N N1) - [* >Ha -Ha #H1 destruct (H1) #Ha - %{N1} (* CR-term *) % [@R_to_star //|@R_to_star @(riota … Ha)] - |#Ha1 %{M2} (* CR-term *) % [// | @R_to_star @riota length_map >length_map //] #n #M0 - cases (true_or_false (leb (|enum (FinSet_of_FType O D ty)|) n)) #Hcase - [>nth_to_default [2:>length_map @(leb_true_to_le … Hcase)] - >nth_to_default [2:>length_map @(leb_true_to_le … Hcase)] // - |cut (n < |enum (FinSet_of_FType O D ty)|) - [@not_le_to_lt @leb_false_to_not_le @Hcase] #Hlt - cut (∃a:FinSet_of_FType O D ty.True) - [lapply Hlt lapply (enum_complete (FinSet_of_FType O D ty)) - cases (enum (FinSet_of_FType O D ty)) - [#_ normalize #H @False_ind @(absurd … H) @lt_to_not_le // - |#a #l #_ #_ %{a} // - ] - ] * #a #_ - >(nth_map ?????? a Hlt) >(nth_map ?????? a Hlt) #_ - @red_star_subst2 // - ] - ] -qed. - -(* we need to proceed by structural induction on the term and then -by inversion on the two redexes. The problem are the moves in a -same subterm, since we need an induction hypothesis, there *) - -lemma local_confluence: ∀O,D,M,M1,M2. red O D M M1 → red O D M M2 → -∃M3. star ? (red O D) M1 M3 ∧ star ? (red O D) M2 M3. -#O #D #M @(T_elim … M) - [#o #a #M1 #M2 #H elim(red_val ????? H) - |#n #M1 #M2 #H elim(red_rel ???? H) - |(* app : this is the interesting case *) - #P #Q #HindP #HindQ - #M1 #M2 #H1 inversion H1 -H1 - [(* right redex is beta *) - #ty #Q #N #Hc #HM >HM -HM #HM1 >HM1 - HM1 #Hl inversion Hl - [#ty1 #Q1 #N1 #Hc1 #H1 destruct (H1) #H_ - %{(subst O D Q1 0 N1)} (* CR-term *) /2/ - |#ty #v #a #M0 #_ #H1 destruct (H1) (* vacuous *) - |#M0 #M10 #N0 #redM0 #_ #H1 destruct (H1) #_ cases (red_lambda … redM0) - [* #Q1 * #redQ #HM10 >HM10 - %{(subst O D Q1 0 N0)} (* CR-term *) % - [@red_star_subst2 //|@R_to_star @rbeta @Hc] - |#HM1 >HM1 @critical - ] - |#M0 #N0 #N1 #redN0N1 #_ #H1 destruct (H1) #HM2 - %{(subst O D Q 0 N1)} (* CR-term *) - %[@red_star_subst @R_to_star //|@R_to_star @rbeta @(red_closed … Hc) //] - |#ty1 #N0 #N1 #_ #_ #H1 destruct (H1) (* vacuous *) - |#ty1 #M0 #H1 destruct (H1) (* vacuous *) - |#ty1 #N0 #N1 #v #v1 #_ #_ #H1 destruct (H1) (* vacuous *) - ] - |(* right redex is iota *)#ty #v #a #M3 #Ha #_ #_ #Hl inversion Hl - [#P1 #M1 #N1 #_ #H1 destruct (H1) (* vacuous *) - |#ty1 #v1 #a1 #M4 #Ha1 #H1 destruct (H1) -H1 #HM4 >(inj_to_T … e0) in Ha; - >Ha1 #H1 destruct (H1) %{M3} (* CR-term *) /2/ - |#M0 #M10 #N0 #redM0 #_ #H1 destruct (H1) #HM2 @(critical2 … redM0 Hl Ha) - |#M0 #N0 #N1 #redN0N1 #_ #H1 destruct (H1) elim (normal_to_T … redN0N1) - |#ty1 #N0 #N1 #_ #_ #H1 destruct (H1) (* vacuous *) - |#ty1 #M0 #H1 destruct (H1) (* vacuous *) - |#ty1 #N0 #N1 #v #v1 #_ #_ #H1 destruct (H1) (* vacuous *) - ] - |(* right redex is appl *)#M3 #M4 #N #redM3M4 #_ #H1 destruct (H1) #_ - #Hl inversion Hl - [#ty1 #M1 #N1 #Hc #H1 destruct (H1) #HM2 lapply (red_lambda … redM3M4) * - [* #M3 * #H1 #H2 >H2 %{(subst O D M3 0 N1)} % - [@R_to_star @rbeta @Hc|@red_star_subst2 // ] - |#H >H -H lapply (critical O D ty1 M1 N1) * #M3 * #H1 #H2 - %{M3} /2/ - ] - |#ty1 #v1 #a1 #M4 #Ha1 #H1 #H2 destruct - lapply (critical2 … redM3M4 Hl Ha1) * #M3 * #H1 #H2 %{M3} /2/ - |#M0 #M10 #N0 #redM0 #_ #H1 destruct (H1) #HM2 - lapply (HindP … redM0 redM3M4) * #M3 * #H1 #H2 - %{(App O D M3 N0)} (* CR-term *) % [@star_red_appl //|@star_red_appl //] - |#M0 #N0 #N1 #redN0N1 #_ #H1 destruct (H1) #_ - %{(App O D M4 N1)} % @R_to_star [@rappr //|@rappl //] - |#ty1 #N0 #N1 #_ #_ #H1 destruct (H1) (* vacuous *) - |#ty1 #M0 #H1 destruct (H1) (* vacuous *) - |#ty1 #N0 #N1 #v #v1 #_ #_ #H1 destruct (H1) (* vacuous *) - ] - |(* right redex is appr *)#M3 #N #N1 #redN #_ #H1 destruct (H1) #_ - #Hl inversion Hl - [#ty1 #M0 #N0 #Hc #H1 destruct (H1) #HM2 - %{(subst O D M0 0 N1)} (* CR-term *) % - [@R_to_star @rbeta @(red_closed … Hc) //|@red_star_subst @R_to_star // ] - |#ty1 #v1 #a1 #M4 #Ha1 #H1 #H2 destruct (H1) elim (normal_to_T … redN) - |#M0 #M10 #N0 #redM0 #_ #H1 destruct (H1) #HM2 - %{(App O D M10 N1)} (* CR-term *) % @R_to_star [@rappl //|@rappr //] - |#M0 #N0 #N10 #redN0 #_ #H1 destruct (H1) #_ - lapply (HindQ … redN0 redN) * #M3 * #H1 #H2 - %{(App O D M0 M3)} (* CR-term *) % [@star_red_appr //|@star_red_appr //] - |#ty1 #N0 #N1 #_ #_ #H1 destruct (H1) (* vacuous *) - |#ty1 #M0 #H1 destruct (H1) (* vacuous *) - |#ty1 #N0 #N1 #v #v1 #_ #_ #H1 destruct (H1) (* vacuous *) - ] - |(* right redex is rlam *) #ty #N0 #N1 #_ #_ #H1 destruct (H1) (* vacuous *) - |(* right redex is rmem *) #ty #M0 #H1 destruct (H1) (* vacuous *) - |(* right redex is vec *) #ty #N #N1 #v #v1 #_ #_ - #H1 destruct (H1) (* vacuous *) - ] - |#ty #M1 #Hind #M2 #M3 #H1 #H2 (* this case is not trivial any more *) - lapply (red_lambda … H1) * - [* #M4 * #H3 #H4 >H4 lapply (red_lambda … H2) * - [* #M5 * #H5 #H6 >H6 lapply(Hind … H3 H5) * #M6 * #H7 #H8 - %{(Lambda O D ty M6)} (* CR-term *) % @star_red_lambda // - |#H5 >H5 @critical3 // - ] - |#HM2 >HM2 lapply (red_lambda … H2) * - [* #M4 * #Hred #HM3 >HM3 lapply (critical3 … ty ?? Hred) * #M5 - * #H3 #H4 %{M5} (* CR-term *) % // - |#HM3 >HM3 %{M3} (* CR-term *) % // - ] - ] - |#ty #v1 #Hind #M1 #M2 #H1 #H2 - lapply (red_vec … H1) * #N11 * #N12 * #v11 * #v12 * * #redN11 #Hv1 #HM1 - lapply (red_vec … H2) * #N21* #N22 * #v21 * #v22 * * #redN21 #Hv2 #HM2 - >Hv1 in Hv2; #Hvv lapply (compare_append … Hvv) -Hvv * - (* we must proceed by cases on the list *) * normalize - [(* N11 = N21 *) * - [>append_nil * #Hl1 #Hl2 destruct lapply(Hind N11 … redN11 redN21) - [@mem_append_l2 %1 //] - * #M3 * #HM31 #HM32 - %{(Vec O D ty (v21@M3::v12))} (* CR-term *) - % [@star_red_vec //|@star_red_vec //] - |>append_nil * #Hl1 #Hl2 destruct lapply(Hind N21 … redN21 redN11) - [@mem_append_l2 %1 //] - * #M3 * #HM31 #HM32 - %{(Vec O D ty (v11@M3::v22))} (* CR-term *) - % [@star_red_vec //|@star_red_vec //] - ] - |(* N11 ≠ N21 *) -Hind #P #l * - [* #Hv11 #Hv22 destruct - %{((Vec O D ty ((v21@N22::l)@N12::v12)))} (* CR-term *) % @R_to_star - [>associative_append >associative_append normalize @rvec // - |>append_cons append_cons associative_append >associative_append normalize @rvec // - ] - ] - ] - ] -qed. - - - - diff --git a/matita/matita/lib/finite_lambda/reduction.ma b/matita/matita/lib/finite_lambda/reduction.ma deleted file mode 100644 index 98c56e10a..000000000 --- a/matita/matita/lib/finite_lambda/reduction.ma +++ /dev/null @@ -1,308 +0,0 @@ -(* - ||M|| This file is part of HELM, an Hypertextual, Electronic - ||A|| Library of Mathematics, developed at the Computer Science - ||T|| Department of the University of Bologna, Italy. - ||I|| - ||T|| - ||A|| This file is distributed under the terms of the - \ / GNU General Public License Version 2 - \ / - V_______________________________________________________________ *) - -include "finite_lambda/terms_and_types.ma". - -(* some auxiliary lemmas *) - -lemma nth_to_default: ∀A,l,n,d. - |l| ≤ n → nth n A l d = d. -#A #l elim l [//] #a #tl #Hind #n cases n - [#d normalize #H @False_ind @(absurd … H) @lt_to_not_le // - |#m #d normalize #H @Hind @le_S_S_to_le @H - ] -qed. - -lemma mem_nth: ∀A,l,n,d. - n < |l| → mem ? (nth n A l d) l. -#A #l elim l - [#n #d normalize #H @False_ind @(absurd … H) @lt_to_not_le // - |#a #tl #Hind * normalize - [#_ #_ %1 //| #m #d #HSS %2 @Hind @le_S_S_to_le @HSS] - ] -qed. - -lemma nth_map: ∀A,B,l,f,n,d1,d2. - n < |l| → nth n B (map … f l) d1 = f (nth n A l d2). -#n #B #l #f elim l - [#m #d1 #d2 normalize #H @False_ind @(absurd … H) @lt_to_not_le // - |#a #tl #Hind #m #d1 #d2 cases m normalize // - #m1 #H @Hind @le_S_S_to_le @H - ] -qed. - - - -(* end of auxiliary lemmas *) - -let rec to_T O D ty on ty: FinSet_of_FType O D ty → T O D ≝ - match ty return (λty.FinSet_of_FType O D ty → T O D) with - [atom o ⇒ λa.Val O D o a - |arrow ty1 ty2 ⇒ λa:FinFun ??.Vec O D ty1 - (map ((FinSet_of_FType O D ty1)×(FinSet_of_FType O D ty2)) - (T O D) (λp.to_T O D ty2 (snd … p)) (pi1 … a)) - ] -. - -lemma is_closed_to_T: ∀O,D,ty,a. is_closed O D 0 (to_T O D ty a). -#O #D #ty elim ty // -#ty1 #ty2 #Hind1 #Hind2 #a normalize @cvec #m #Hmem -lapply (mem_map ????? Hmem) * #a1 * #H1 #H2

H //] -Hl1 -Hl2 - lapply e0 -e0 lapply l2 -l2 elim l1 - [#l2 cases l2 normalize [// |#a1 #tl1 #H destruct] - |#a1 #tl1 #Hind #l2 cases l2 - [normalize #H destruct - |#a2 #tl2 normalize #H @eq_f2 - [@Hind2 *) - -let rec assoc (A:FinSet) (B:Type[0]) (a:A) l1 l2 on l1 : option B ≝ - match l1 with - [ nil ⇒ None ? - | cons hd1 tl1 ⇒ match l2 with - [ nil ⇒ None ? - | cons hd2 tl2 ⇒ if a==hd1 then Some ? hd2 else assoc A B a tl1 tl2 - ] - ]. - -lemma same_assoc: ∀A,B,a,l1,v1,v2,N,N1. - assoc A B a l1 (v1@N::v2) = Some ? N ∧ assoc A B a l1 (v1@N1::v2) = Some ? N1 - ∨ assoc A B a l1 (v1@N::v2) = assoc A B a l1 (v1@N1::v2). -#A #B #a #l1 #v1 #v2 #N #N1 lapply v1 -v1 elim l1 - [#v1 %2 // |#hd #tl #Hind * normalize cases (a==hd) normalize /3/] -qed. - -lemma assoc_to_mem: ∀A,B,a,l1,l2,b. - assoc A B a l1 l2 = Some ? b → mem ? b l2. -#A #B #a #l1 elim l1 - [#l2 #b normalize #H destruct - |#hd1 #tl1 #Hind * - [#b normalize #H destruct - |#hd2 #tl2 #b normalize cases (a==hd1) normalize - [#H %1 destruct //|#H %2 @Hind @H] - ] - ] -qed. - -lemma assoc_to_mem2: ∀A,B,a,l1,l2,b. - assoc A B a l1 l2 = Some ? b → ∃l21,l22.l2=l21@b::l22. -#A #B #a #l1 elim l1 - [#l2 #b normalize #H destruct - |#hd1 #tl1 #Hind * - [#b normalize #H destruct - |#hd2 #tl2 #b normalize cases (a==hd1) normalize - [#H %{[]} %{tl2} destruct // - |#H lapply (Hind … H) * #la * #lb #H1 - %{(hd2::la)} %{lb} >H1 //] - ] - ] -qed. - -lemma assoc_map: ∀A,B,C,a,l1,l2,f,b. - assoc A B a l1 l2 = Some ? b → assoc A C a l1 (map ?? f l2) = Some ? (f b). -#A #B #C #a #l1 elim l1 - [#l2 #f #b normalize #H destruct - |#hd1 #tl1 #Hind * - [#f #b normalize #H destruct - |#hd2 #tl2 #f #b normalize cases (a==hd1) normalize - [#H destruct // |#H @(Hind … H)] - ] - ] -qed. - -(*************************** One step reduction *******************************) - -inductive red (O:Type[0]) (D:O→FinSet) : T O D →T O D → Prop ≝ - | (* we only allow beta on closed arguments *) - rbeta: ∀P,M,N. is_closed O D 0 N → - red O D (App O D (Lambda O D P M) N) (subst O D M 0 N) - | riota: ∀ty,v,a,M. - assoc ?? a (enum (FinSet_of_FType O D ty)) v = Some ? M → - red O D (App O D (Vec O D ty v) (to_T O D ty a)) M - | rappl: ∀M,M1,N. red O D M M1 → red O D (App O D M N) (App O D M1 N) - | rappr: ∀M,N,N1. red O D N N1 → red O D (App O D M N) (App O D M N1) - | rlam: ∀ty,N,N1. red O D N N1 → red O D (Lambda O D ty N) (Lambda O D ty N1) - | rmem: ∀ty,M. red O D (Lambda O D ty M) - (Vec O D ty (map ?? (λa. subst O D M 0 (to_T O D ty a)) - (enum (FinSet_of_FType O D ty)))) - | rvec: ∀ty,N,N1,v,v1. red O D N N1 → - red O D (Vec O D ty (v@N::v1)) (Vec O D ty (v@N1::v1)). - -(*********************************** inversion ********************************) -lemma red_vec: ∀O,D,ty,v,M. - red O D (Vec O D ty v) M → ∃N,N1,v1,v2. - red O D N N1 ∧ v = v1@N::v2 ∧ M = Vec O D ty (v1@N1::v2). -#O #D #ty #v #M #Hred inversion Hred - [#ty1 #M0 #N #Hc #H destruct - |#ty1 #v1 #a #M0 #_ #H destruct - |#M0 #M1 #N #_ #_ #H destruct - |#M0 #M1 #N #_ #_ #H destruct - |#ty1 #M #M1 #_ #_ #H destruct - |#ty1 #M0 #H destruct - |#ty1 #N #N1 #v1 #v2 #Hred1 #_ #H destruct #_ %{N} %{N1} %{v1} %{v2} /3/ - ] -qed. - -lemma red_lambda: ∀O,D,ty,M,N. - red O D (Lambda O D ty M) N → - (∃M1. red O D M M1 ∧ N = (Lambda O D ty M1)) ∨ - N = Vec O D ty (map ?? (λa. subst O D M 0 (to_T O D ty a)) - (enum (FinSet_of_FType O D ty))). -#O #D #ty #M #N #Hred inversion Hred - [#ty1 #M0 #N #Hc #H destruct - |#ty1 #v1 #a #M0 #_ #H destruct - |#M0 #M1 #N #_ #_ #H destruct - |#M0 #M1 #N #_ #_ #H destruct - |#ty1 #P #P1 #redP #_ #H #H1 destruct %1 %{P1} % // - |#ty1 #M0 #H destruct #_ %2 // - |#ty1 #N #N1 #v1 #v2 #Hred1 #_ #H destruct - ] -qed. - -lemma red_val: ∀O,D,ty,a,N. - red O D (Val O D ty a) N → False. -#O #D #ty #M #N #Hred inversion Hred - [#ty1 #M0 #N #Hc #H destruct - |#ty1 #v1 #a #M0 #_ #H destruct - |#M0 #M1 #N #_ #_ #H destruct - |#M0 #M1 #N #_ #_ #H destruct - |#ty1 #N1 #N2 #_ #_ #H destruct - |#ty1 #M0 #H destruct #_ - |#ty1 #N #N1 #v1 #v2 #Hred1 #_ #H destruct - ] -qed. - -lemma red_rel: ∀O,D,n,N. - red O D (Rel O D n) N → False. -#O #D #n #N #Hred inversion Hred - [#ty1 #M0 #N #Hc #H destruct - |#ty1 #v1 #a #M0 #_ #H destruct - |#M0 #M1 #N #_ #_ #H destruct - |#M0 #M1 #N #_ #_ #H destruct - |#ty1 #N1 #N2 #_ #_ #H destruct - |#ty1 #M0 #H destruct #_ - |#ty1 #N #N1 #v1 #v2 #Hred1 #_ #H destruct - ] -qed. - -(*************************** multi step reduction *****************************) -lemma star_red_appl: ∀O,D,M,M1,N. star ? (red O D) M M1 → - star ? (red O D) (App O D M N) (App O D M1 N). -#O #D #M #N #N1 #H elim H // -#P #Q #Hind #HPQ #Happ %1[|@Happ] @rappl @HPQ -qed. - -lemma star_red_appr: ∀O,D,M,N,N1. star ? (red O D) N N1 → - star ? (red O D) (App O D M N) (App O D M N1). -#O #D #M #N #N1 #H elim H // -#P #Q #Hind #HPQ #Happ %1[|@Happ] @rappr @HPQ -qed. - -lemma star_red_vec: ∀O,D,ty,N,N1,v1,v2. star ? (red O D) N N1 → - star ? (red O D) (Vec O D ty (v1@N::v2)) (Vec O D ty (v1@N1::v2)). -#O #D #ty #N #N1 #v1 #v2 #H elim H // -#P #Q #Hind #HPQ #Hvec %1[|@Hvec] @rvec @HPQ -qed. - -lemma star_red_vec1: ∀O,D,ty,v1,v2,v. |v1| = |v2| → - (∀n,M. n < |v1| → star ? (red O D) (nth n ? v1 M) (nth n ? v2 M)) → - star ? (red O D) (Vec O D ty (v@v1)) (Vec O D ty (v@v2)). -#O #D #ty #v1 elim v1 - [#v2 #v normalize #Hv2 >(lenght_to_nil … (sym_eq … Hv2)) normalize // - |#N1 #tl1 #Hind * [normalize #v #H destruct] #N2 #tl2 #v normalize #HS - #H @(trans_star … (Vec O D ty (v@N2::tl1))) - [@star_red_vec @(H 0 N1) @le_S_S // - |>append_cons >(append_cons ??? tl2) @(Hind… (injective_S … HS)) - #n #M #H1 @(H (S n)) @le_S_S @H1 - ] - ] -qed. - -lemma star_red_vec2: ∀O,D,ty,v1,v2. |v1| = |v2| → - (∀n,M. n < |v1| → star ? (red O D) (nth n ? v1 M) (nth n ? v2 M)) → - star ? (red O D) (Vec O D ty v1) (Vec O D ty v2). -#O #D #ty #v1 #v2 @(star_red_vec1 … [ ]) -qed. - -lemma star_red_lambda: ∀O,D,ty,N,N1. star ? (red O D) N N1 → - star ? (red O D) (Lambda O D ty N) (Lambda O D ty N1). -#O #D #ty #N #N1 #H elim H // -#P #Q #Hind #HPQ #Hlam %1[|@Hlam] @rlam @HPQ -qed. - -(************************ reduction and substitution **************************) - -lemma red_star_subst : ∀O,D,M,N,N1,i. - star ? (red O D) N N1 → star ? (red O D) (subst O D M i N) (subst O D M i N1). -#O #D #M #N #N1 #i #Hred lapply i -i @(T_elim … M) normalize - [#o #a #i // - |#i #n cases (leb n i) normalize // cases (eqb n i) normalize // - |#P #Q #HindP #HindQ #n normalize - @(trans_star … (App O D (subst O D P n N1) (subst O D Q n N))) - [@star_red_appl @HindP |@star_red_appr @HindQ] - |#ty #P #HindP #i @star_red_lambda @HindP - |#ty #v #Hindv #i @star_red_vec2 [>length_map >length_map //] - #j #Q inversion v [#_ normalize //] #a #tl #_ #Hv - cases (true_or_false (leb (S j) (|a::tl|))) #Hcase - [lapply (leb_true_to_le … Hcase) -Hcase #Hcase - >(nth_map ?????? a Hcase) >(nth_map ?????? a Hcase) #_ @Hindv >Hv @mem_nth // - |>nth_to_default - [2:>length_map @le_S_S_to_le @not_le_to_lt @leb_false_to_not_le //] - >nth_to_default - [2:>length_map @le_S_S_to_le @not_le_to_lt @leb_false_to_not_le //] // - ] - ] -qed. - -lemma red_star_subst2 : ∀O,D,M,M1,N,i. is_closed O D 0 N → - red O D M M1 → star ? (red O D) (subst O D M i N) (subst O D M1 i N). -#O #D #M #M1 #N #i #HNc #Hred lapply i -i elim Hred - [#ty #P #Q #HQc #i normalize @starl_to_star @sstepl - [|@rbeta >(subst_closed … HQc) //] >(subst_closed … HQc) // - lapply (subst_lemma ?? P ?? i 0 (is_closed_mono … HQc) HNc) // - (subst_closed … (le_O_n …)) // - @R_to_star @riota @assoc_map @HP - |#P #P1 #Q #Hred #Hind #i normalize @star_red_appl @Hind - |#P #P1 #Q #Hred #Hind #i normalize @star_red_appr @Hind - |#ty #P #P1 #Hred #Hind #i normalize @star_red_lambda @Hind - |#ty #P #i normalize @starl_to_star @sstepl [|@rmem] - @star_to_starl @star_red_vec2 [>length_map >length_map >length_map //] - #n #Q >length_map #H - cut (∃a:(FinSet_of_FType O D ty).True) - [lapply H -H lapply (enum_complete (FinSet_of_FType O D ty)) - cases (enum (FinSet_of_FType O D ty)) - [#x normalize #H @False_ind @(absurd … H) @lt_to_not_le // - |#x #l #_ #_ %{x} // - ] - ] * #a #_ - >(nth_map ?????? a H) >(nth_map ?????? Q) [2:>length_map @H] - >(nth_map ?????? a H) - lapply (subst_lemma O D P (to_T O D ty - (nth n (FinSet_of_FType O D ty) (enum (FinSet_of_FType O D ty)) a)) - N i 0 (is_closed_mono … (is_closed_to_T …)) HNc) // H1 - Hx @Hind2 normalize // - |@Hind1 #N #H @Hind2 @(lt_to_le_to_lt … H) normalize // - ] - ] - ] -qed. - -(* since we only consider beta reduction with closed arguments we could avoid -lifting. We define it for the sake of generality *) - -(* arguments: k is the nesting depth (starts from 0), p is the lift -let rec lift O D t k p on t ≝ - match t with - [ Val o a ⇒ Val O D o a - | Rel n ⇒ if (leb k n) then Rel O D (n+p) else Rel O D n - | App m n ⇒ App O D (lift O D m k p) (lift O D n k p) - | Lambda Ty n ⇒ Lambda O D Ty (lift O D n (S k) p) - | Vec Ty v ⇒ Vec O D Ty (map ?? (λx. lift O D x k p) v) - ]. - -notation "↑ ^ n ( M )" non associative with precedence 40 for @{'Lift 0 $n $M}. -notation "↑ _ k ^ n ( M )" non associative with precedence 40 for @{'Lift $n $k $M}. - -interpretation "Lift" 'Lift n k M = (lift ?? M k n). - -let rec subst O D t k s on t ≝ - match t with - [ Val o a ⇒ Val O D o a - | Rel n ⇒ if (leb k n) then - (if (eqb k n) then lift O D s 0 n else Rel O D (n-1)) - else(Rel O D n) - | App m n ⇒ App O D (subst O D m k s) (subst O D n k s) - | Lambda T n ⇒ Lambda O D T (subst O D n (S k) s) - | Vec T v ⇒ Vec O D T (map ?? (λx. subst O D x k s) v) - ]. -*) - -(* simplified version of subst, assuming the argument s is closed *) - -let rec subst O D t k s on t ≝ - match t with - [ Val o a ⇒ Val O D o a - | Rel n ⇒ if (leb k n) then - (if (eqb k n) then (* lift O D s 0 n*) s else Rel O D (n-1)) - else(Rel O D n) - | App m n ⇒ App O D (subst O D m k s) (subst O D n k s) - | Lambda T n ⇒ Lambda O D T (subst O D n (S k) s) - | Vec T v ⇒ Vec O D T (map ?? (λx. subst O D x k s) v) - ]. -(* notation "hvbox(M break [ k ≝ N ])" - non associative with precedence 90 - for @{'Subst1 $M $k $N}. *) - -interpretation "Subst" 'Subst1 M k N = (subst M k N). - -(* -lemma subst_rel1: ∀O,D,A.∀k,i. i < k → - (Rel O D i) [k ≝ A] = Rel O D i. -#A #k #i normalize #ltik >(lt_to_leb_false … ltik) // -qed. - -lemma subst_rel2: ∀O,D, A.∀k. - (Rel k) [k ≝ A] = lift A 0 k. -#A #k normalize >(le_to_leb_true k k) // >(eq_to_eqb_true … (refl …)) // -qed. - -lemma subst_rel3: ∀A.∀k,i. k < i → - (Rel i) [k ≝ A] = Rel (i-1). -#A #k #i normalize #ltik >(le_to_leb_true k i) /2/ ->(not_eq_to_eqb_false k i) // @lt_to_not_eq // -qed. *) - - -(* closed terms ???? -let rec closed_k O D (t: T O D) k on t ≝ - match t with - [ Val o a ⇒ True - | Rel n ⇒ n < k - | App m n ⇒ (closed_k O D m k) ∧ (closed_k O D n k) - | Lambda T n ⇒ closed_k O D n (k+1) - | Vec T v ⇒ closed_list O D v k - ] - -and closed_list O D (l: list (T O D)) k on l ≝ - match l with - [ nil ⇒ True - | cons hd tl ⇒ closed_k O D hd k ∧ closed_list O D tl k - ] -. *) - -inductive is_closed (O:Type[0]) (D:O→FinSet): nat → T O D → Prop ≝ -| cval : ∀k,o,a.is_closed O D k (Val O D o a) -| crel : ∀k,n. n < k → is_closed O D k (Rel O D n) -| capp : ∀k,m,n. is_closed O D k m → is_closed O D k n → - is_closed O D k (App O D m n) -| clam : ∀T,k,m. is_closed O D (S k) m → is_closed O D k (Lambda O D T m) -| cvec: ∀T,k,v. (∀m. mem ? m v → is_closed O D k m) → - is_closed O D k (Vec O D T v). - -lemma is_closed_rel: ∀O,D,n,k. - is_closed O D k (Rel O D n) → n < k. -#O #D #n #k #H inversion H - [#k0 #o #a #eqk #H destruct - |#k0 #n0 #ltn0 #eqk #H destruct // - |#k0 #M #N #_ #_ #_ #_ #_ #H destruct - |#T #k0 #M #_ #_ #_ #H destruct - |#T #k0 #v #_ #_ #_ #H destruct - ] -qed. - -lemma is_closed_app: ∀O,D,k,M, N. - is_closed O D k (App O D M N) → is_closed O D k M ∧ is_closed O D k N. -#O #D #k #M #N #H inversion H - [#k0 #o #a #eqk #H destruct - |#k0 #n0 #ltn0 #eqk #H destruct - |#k0 #M1 #N1 #HM #HN #_ #_ #_ #H1 destruct % // - |#T #k0 #M #_ #_ #_ #H destruct - |#T #k0 #v #_ #_ #_ #H destruct - ] -qed. - -lemma is_closed_lam: ∀O,D,k,ty,M. - is_closed O D k (Lambda O D ty M) → is_closed O D (S k) M. -#O #D #k #ty #M #H inversion H - [#k0 #o #a #eqk #H destruct - |#k0 #n0 #ltn0 #eqk #H destruct - |#k0 #M1 #N1 #HM #HN #_ #_ #_ #H1 destruct - |#T #k0 #M1 #HM1 #_ #_ #H1 destruct // - |#T #k0 #v #_ #_ #_ #H destruct - ] -qed. - -lemma is_closed_vec: ∀O,D,k,ty,v. - is_closed O D k (Vec O D ty v) → ∀m. mem ? m v → is_closed O D k m. -#O #D #k #ty #M #H inversion H - [#k0 #o #a #eqk #H destruct - |#k0 #n0 #ltn0 #eqk #H destruct - |#k0 #M1 #N1 #HM #HN #_ #_ #_ #H1 destruct - |#T #k0 #M1 #HM1 #_ #_ #H1 destruct - |#T #k0 #v #Hv #_ #_ #H1 destruct @Hv - ] -qed. - -lemma is_closed_S: ∀O,D,M,m. - is_closed O D m M → is_closed O D (S m) M. -#O #D #M #m #H elim H // - [#k #n0 #Hlt @crel @le_S // - |#k #P #Q #HP #HC #H1 #H2 @capp // - |#ty #k #P #HP #H1 @clam // - |#ty #k #v #Hind #Hv @cvec @Hv - ] -qed. - -lemma is_closed_mono: ∀O,D,M,m,n. m ≤ n → - is_closed O D m M → is_closed O D n M. -#O #D #M #m #n #lemn elim lemn // #i #j #H #H1 @is_closed_S @H @H1 -qed. - - -(*** properties of lift and subst ***) - -(* -lemma lift_0: ∀O,D.∀t:T O D.∀k. lift O D t k 0 = t. -#O #D #t @(T_elim … t) normalize // - [#n #k cases (leb k n) normalize // - |#o #v #Hind #k @eq_f lapply Hind -Hind elim v // - #hd #tl #Hind #Hind1 normalize @eq_f2 - [@Hind1 %1 //|@Hind #x #Hx @Hind1 %2 //] - ] -qed. - -lemma lift_closed: ∀O,D.∀t:T O D.∀k,p. - is_closed O D k t → lift O D t k p = t. -#O #D #t @(T_elim … t) normalize // - [#n #k #p #H >(not_le_to_leb_false … (lt_to_not_le … (is_closed_rel … H))) // - |#M #N #HindM #HindN #k #p #H lapply (is_closed_app … H) * #HcM #HcN - >(HindM … HcM) >(HindN … HcN) // - |#ty #M #HindM #k #p #H lapply (is_closed_lam … H) -H #H >(HindM … H) // - |#ty #v #HindM #k #p #H lapply (is_closed_vec … H) -H #H @eq_f - cut (∀m. mem ? m v → lift O D m k p = m) - [#m #Hmem @HindM [@Hmem | @H @Hmem]] -HindM - elim v // #a #tl #Hind #H1 normalize @eq_f2 - [@H1 %1 //|@Hind #m #Hmem @H1 %2 @Hmem] - ] -qed. - -*) - -lemma subst_closed: ∀O,D,M,N,k,i. k ≤ i → - is_closed O D k M → subst O D M i N = M. -#O #D #M @(T_elim … M) - [#o #a normalize // - |#n #N #k #j #Hlt #Hc lapply (is_closed_rel … Hc) #Hnk normalize - >not_le_to_leb_false [2:@lt_to_not_le @(lt_to_le_to_lt … Hnk Hlt)] // - |#P #Q #HindP #HindQ #N #k #i #ltki #Hc lapply (is_closed_app … Hc) * - #HcP #HcQ normalize >(HindP … ltki HcP) >(HindQ … ltki HcQ) // - |#ty #P #HindP #N #k #i #ltki #Hc lapply (is_closed_lam … Hc) - #HcP normalize >(HindP … HcP) // @le_S_S @ltki - |#ty #v #Hindv #N #k #i #ltki #Hc lapply (is_closed_vec … Hc) - #Hcv normalize @eq_f - cut (∀m:T O D.mem (T O D) m v→ subst O D m i N=m) - [#m #Hmem @(Hindv … Hmem N … ltki) @Hcv @Hmem] - elim v // #a #tl #Hind #H normalize @eq_f2 - [@H %1 //| @Hind #Hmem #Htl @H %2 @Htl] - ] -qed. - -lemma subst_lemma: ∀O,D,A,B,C,k,i. is_closed O D k B → is_closed O D i C → - subst O D (subst O D A i B) (k+i) C = - subst O D (subst O D A (k+S i) C) i B. -#O #D #A #B #C #k @(T_elim … A) normalize - [// - |#n #i #HBc #HCc @(leb_elim i n) #Hle - [@(eqb_elim i n) #eqni - [(lt_to_leb_false (k+(S i)) i) // normalize - >(subst_closed … HBc) // >le_to_leb_true // >eq_to_eqb_true // - |(cut (i < n)) - [cases (le_to_or_lt_eq … Hle) // #eqin @False_ind /2/] #ltin - (cut (0 < n)) [@(le_to_lt_to_lt … ltin) //] #posn - normalize @(leb_elim (k+i) (n-1)) #nk - [@(eqb_elim (k+i) (n-1)) #H normalize - [cut (k+(S i) = n); [/2 by S_pred/] #H1 - >(le_to_leb_true (k+(S i)) n) /2/ - >(eq_to_eqb_true … H1) normalize >(subst_closed … HCc) // - |(cut (k+i < n-1)) [@not_eq_to_le_to_lt; //] #Hlt - >(le_to_leb_true (k+(S i)) n) normalize - [>(not_eq_to_eqb_false (k+(S i)) n) normalize - [>le_to_leb_true [2:@lt_to_le @(le_to_lt_to_lt … Hlt) //] - >not_eq_to_eqb_false // @lt_to_not_eq @(le_to_lt_to_lt … Hlt) // - |@(not_to_not … H) #Hn /2 by plus_minus/ - ] - |(not_le_to_leb_false (k+(S i)) n) normalize - [>(le_to_leb_true … Hle) >(not_eq_to_eqb_false … eqni) // - |@(not_to_not … nk) #H @le_plus_to_minus_r // - ] - ] - ] - |(cut (n < k+i)) [@(lt_to_le_to_lt ? i) /2 by not_le_to_lt/] #ltn - >not_le_to_leb_false [2: @lt_to_not_le @(transitive_lt …ltn) //] normalize - >not_le_to_leb_false [2: @lt_to_not_le //] normalize - >(not_le_to_leb_false … Hle) // - ] - |#M #N #HindM #HindN #i #HBC #HCc @eq_f2 [@HindM // |@HindN //] - |#ty #M #HindM #i #HBC #HCc @eq_f >plus_n_Sm >plus_n_Sm @HindM // - @is_closed_S // - |#ty #v #Hindv #i #HBC #HCc @eq_f - cut (∀m.mem ? m v → subst O D (subst O D m i B) (k+i) C = - subst O D (subst O D m (k+S i) C) i B) - [#m #Hmem @Hindv //] -Hindv elim v normalize [//] - #a #tl #Hind #H @eq_f2 [@H %1 // | @Hind #m #Hmem @H %2 //] - ] -qed. - - diff --git a/matita/matita/lib/finite_lambda/typing.ma b/matita/matita/lib/finite_lambda/typing.ma deleted file mode 100644 index 791e27d66..000000000 --- a/matita/matita/lib/finite_lambda/typing.ma +++ /dev/null @@ -1,229 +0,0 @@ -(* - ||M|| This file is part of HELM, an Hypertextual, Electronic - ||A|| Library of Mathematics, developed at the Computer Science - ||T|| Department of the University of Bologna, Italy. - ||I|| - ||T|| - ||A|| This file is distributed under the terms of the - \ / GNU General Public License Version 2 - \ / - V_______________________________________________________________ *) - -include "finite_lambda/reduction.ma". - - -(****************************************************************) - -inductive TJ (O: Type[0]) (D:O → FinSet): list (FType O) → T O D → FType O → Prop ≝ - | tval: ∀G,o,a. TJ O D G (Val O D o a) (atom O o) - | trel: ∀G1,ty,G2,n. length ? G1 = n → TJ O D (G1@ty::G2) (Rel O D n) ty - | tapp: ∀G,M,N,ty1,ty2. TJ O D G M (arrow O ty1 ty2) → TJ O D G N ty1 → - TJ O D G (App O D M N) ty2 - | tlambda: ∀G,M,ty1,ty2. TJ O D (ty1::G) M ty2 → - TJ O D G (Lambda O D ty1 M) (arrow O ty1 ty2) - | tvec: ∀G,v,ty1,ty2. - (|v| = |enum (FinSet_of_FType O D ty1)|) → - (∀M. mem ? M v → TJ O D G M ty2) → - TJ O D G (Vec O D ty1 v) (arrow O ty1 ty2). - -lemma wt_to_T: ∀O,D,G,ty,a.TJ O D G (to_T O D ty a) ty. -#O #D #G #ty elim ty - [#o #a normalize @tval - |#ty1 #ty2 #Hind1 #Hind2 normalize * #v #Hv @tvec - [length_map >length_map // - |#M elim v - [normalize @False_ind |#a #v1 #Hind3 * [#eqM >eqM @Hind2 |@Hind3]] - ] - ] -qed. - -lemma inv_rel: ∀O,D,G,n,ty. - TJ O D G (Rel O D n) ty → ∃G1,G2.|G1|=n∧G=G1@ty::G2. -#O #D #G #n #ty #Hrel inversion Hrel - [#G1 #o #a #_ #H destruct - |#G1 #ty1 #G2 #n1 #H1 #H2 #H3 #H4 destruct %{G1} %{G2} /2/ - |#G1 #M0 #N #ty1 #ty2 #_ #_ #_ #_ #_ #H destruct - |#G1 #M0 #ty4 #ty5 #HM0 #_ #_ #H #H1 destruct - |#G1 #v #ty3 #ty4 #_ #_ #_ #_ #H destruct - ] -qed. - -lemma inv_tlambda: ∀O,D,G,M,ty1,ty2,ty3. - TJ O D G (Lambda O D ty1 M) (arrow O ty2 ty3) → - ty1 = ty2 ∧ TJ O D (ty2::G) M ty3. -#O #D #G #M #ty1 #ty2 #ty3 #Hlam inversion Hlam - [#G1 #o #a #_ #H destruct - |#G1 #ty #G2 #n #_ #_ #H destruct - |#G1 #M0 #N #ty1 #ty2 #_ #_ #_ #_ #_ #H destruct - |#G1 #M0 #ty4 #ty5 #HM0 #_ #_ #H #H1 destruct % // - |#G1 #v #ty3 #ty4 #_ #_ #_ #_ #H destruct - ] -qed. - -lemma inv_tvec: ∀O,D,G,v,ty1,ty2,ty3. - TJ O D G (Vec O D ty1 v) (arrow O ty2 ty3) → - (|v| = |enum (FinSet_of_FType O D ty1)|) ∧ - (∀M. mem ? M v → TJ O D G M ty3). -#O #D #G #v #ty1 #ty2 #ty3 #Hvec inversion Hvec - [#G #o #a #_ #H destruct - |#G1 #ty #G2 #n #_ #_ #H destruct - |#G1 #M0 #N #ty1 #ty2 #_ #_ #_ #_ #_ #H destruct - |#G1 #M0 #ty4 #ty5 #HM0 #_ #_ #H #H1 destruct - |#G1 #v1 #ty4 #ty5 #Hv #Hmem #_ #_ #H #H1 destruct % // @Hmem - ] -qed. - -(* could be generalized *) -lemma weak_rel: ∀O,D,G1,G2,ty1,ty2,n. length ? G1 < n → - TJ O D (G1@G2) (Rel O D n) ty1 → - TJ O D (G1@ty2::G2) (Rel O D (S n)) ty1. -#O #D #G1 #G2 #ty1 #ty2 #n #HG1 #Hrel lapply (inv_rel … Hrel) -* #G3 * #G4 * #H1 #H2 lapply (compare_append … H2) -* #G5 * - [* #H3 @False_ind >H3 in HG1; >length_append >H1 #H4 - @(absurd … H4) @le_to_not_lt // - |* #H3 #H4 >H4 >append_cons length_append >length_append

H3 >length_append normalize - >plus_n_Sm >associative_plus @eq_f // - ] -qed. - -lemma strength_rel: ∀O,D,G1,G2,ty1,ty2,n. length ? G1 < n → - TJ O D (G1@ty2::G2) (Rel O D n) ty1 → - TJ O D (G1@G2) (Rel O D (n-1)) ty1. -#O #D #G1 #G2 #ty1 #ty2 #n #HG1 #Hrel lapply (inv_rel … Hrel) -* #G3 * #G4 * #H1 #H2 lapply (compare_append … H2) -* #G5 * - [* #H3 @False_ind >H3 in HG1; >length_append >H1 #H4 - @(absurd … H4) @le_to_not_lt // - |lapply G5 -G5 * - [>append_nil normalize * #H3 #H4 destruct @False_ind @(absurd … HG1) - @le_to_not_lt // - |#ty3 #G5 * #H3 normalize #H4 destruct (H4) H3 >length_append >length_append normalize H1 >length_append // - |* cases G6 - [>append_nil normalize #H1 @False_ind - @(absurd ? Hlt) @le_to_not_lt H1 // - |#ty1 #G7 #H1 normalize #H2 destruct >associative_append @trel // - ] - ] - |#G #M #N #ty1 #ty2 #HM #HN #HindM #HindN #G1 #G2 #G3 - #Heq #Hc lapply (is_closed_app … Hc) -Hc * #HMc #HNc - @(tapp … (HindM … Heq HMc) (HindN … Heq HNc)) - |#G #M #ty1 #ty2 #HM #HindM #G1 #G2 #G3 #Heq #Hc - lapply (is_closed_lam … Hc) -Hc #HMc - @tlambda @(HindM (ty1::G1) G2) [>Heq // |@HMc] - |#G #v #ty1 #ty2 #Hlen #Hv #Hind #G1 #G2 #G3 #H1 #Hc @tvec - [>length_map // - |#M #Hmem @Hind // lapply (is_closed_vec … Hc) #Hvc @Hvc // - ] - ] -qed. - -lemma nth_spec: ∀A,a,d,l1,l2,n. |l1| = n → nth n A (l1@a::l2) d = a. -#A #a #d #l1 elim l1 normalize - [#l2 #n #Hn H cases (le_to_or_lt_eq … (leb_true_to_le … H)) - [#ltn >(not_eq_to_eqb_false … (lt_to_not_eq … ltn)) normalize - lapply (compare_append … HG) * #G3 * - [* #HG1 #HG2 @(strength_rel … tyN … ltn) HG in ltn; >length_append #ltn @False_ind - @(absurd … ltn) @le_to_not_lt >Hlen // - ] - |#HG21 >(eq_to_eqb_true … HG21) - cut (ty = tyN) - [<(nth_spec ? ty ty ? G2 … Hlen) >HG @nth_spec @HG21] #Hty >Hty - normalize H normalize lapply (compare_append … HG) * #G3 * - [* #H1 @False_ind @(absurd ? Hlen) @sym_not_eq @lt_to_not_eq >H1 - >length_append @(lt_to_le_to_lt n (|G21|)) // @not_le_to_lt - @(leb_false_to_not_le … H) - |cases G3 - [>append_nil * #H1 @False_ind @(absurd ? Hlen)

associative_append @trel // - ] - ] - ] - |#G #M #N #ty1 #ty2 #HM #HN #HindM #HindN #G1 #G2 #N0 #tyN0 #eqG - #HN0 #Hc normalize @(tapp … ty1) - [@(HindM … eqG HN0 Hc) |@(HindN … eqG HN0 Hc)] - |#G #M #ty1 #ty2 #HM #HindM #G1 #G2 #N0 #tyN0 #eqG - #HN0 #Hc normalize @(tlambda … ty1) @(HindM (ty1::G1) … HN0) // >eqG // - |#G #v #ty1 #ty2 #Hlen #Hv #Hind #G1 #G2 #N0 #tyN0 #eqG - #HN0 #Hc normalize @(tvec … ty1) - [>length_map @Hlen - |#M #Hmem lapply (mem_map ????? Hmem) * #a * -Hmem #Hmem #eqM HM1 @tlambda @Hind // - |#HM1 >HM1 @tvec // #N #HN lapply(mem_map ????? HN) - * #a * #mema #eqN H3 @tvec - [H2 >length_append >length_append @eq_f // - |#M2 #Hmem cases (mem_append ???? Hmem) -Hmem #Hmem - [@Hv >H2 @mem_append_l1 // - |cases Hmem - [#HM2 >HM2 -HM2 @(Hind N … H1) >H2 @mem_append_l2 %1 // - |-Hmem #Hmem @Hv >H2 @mem_append_l2 %2 // - ] - ] - ] - ] -qed. - diff --git a/matita/matita/lib/reverse_complexity/big_O.ma b/matita/matita/lib/reverse_complexity/big_O.ma deleted file mode 100644 index 833572ea2..000000000 --- a/matita/matita/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/lib/reverse_complexity/gap.ma b/matita/matita/lib/reverse_complexity/gap.ma deleted file mode 100644 index 2c8b6e938..000000000 --- a/matita/matita/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. - -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/lib/reverse_complexity/speed_clean.ma b/matita/matita/lib/reverse_complexity/speed_clean.ma deleted file mode 100644 index 9b9fecf94..000000000 --- a/matita/matita/lib/reverse_complexity/speed_clean.ma +++ /dev/null @@ -1,1067 +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". - -(************************* 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. - \ No newline at end of file diff --git a/matita/matita/lib/reverse_complexity/speed_def.ma b/matita/matita/lib/reverse_complexity/speed_def.ma deleted file mode 100644 index 6a0ec4a57..000000000 --- a/matita/matita/lib/reverse_complexity/speed_def.ma +++ /dev/null @@ -1,921 +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". - -(************************* 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. - \ No newline at end of file diff --git a/matita/matita/lib/reverse_complexity/speed_new.ma b/matita/matita/lib/reverse_complexity/speed_new.ma deleted file mode 100644 index 6a0ec4a57..000000000 --- a/matita/matita/lib/reverse_complexity/speed_new.ma +++ /dev/null @@ -1,921 +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". - -(************************* 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. - \ No newline at end of file diff --git a/matita/matita/lib/reverse_complexity/toolkit.ma b/matita/matita/lib/reverse_complexity/toolkit.ma deleted file mode 100644 index 11f988c79..000000000 --- a/matita/matita/lib/reverse_complexity/toolkit.ma +++ /dev/null @@ -1,987 +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". - -(************************* 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. - \ No newline at end of file diff --git a/matita/matita/lib/turing/multi_to_mono/multi_to_mono.ma b/matita/matita/lib/turing/multi_to_mono/multi_to_mono.ma deleted file mode 100644 index 0e9bce53e..000000000 --- a/matita/matita/lib/turing/multi_to_mono/multi_to_mono.ma +++ /dev/null @@ -1,397 +0,0 @@ -include "turing/auxiliary_machines1.ma". -include "turing/multi_to_mono/shift_trace_machines.ma". - -(******************************************************************************) -(* mtiL: complete move L for tape i. We reaching the left border of trace i, *) -(* add a blank if there is no more tape, then move the i-trace and finally *) -(* come back to the head position. *) -(******************************************************************************) - -(* we say that a tape is regular if for any trace after the first blank we - only have other blanks *) - -definition all_blanks_in ≝ λsig,l. - ∀x. mem ? x l → x = blank sig. - -definition regular_i ≝ λsig,n.λl:list (multi_sig sig n).λi. - all_blanks_in ? (after_blank ? (trace sig n i l)). - -definition regular_trace ≝ λsig,n,a.λls,rs:list (multi_sig sig n).λi. - Or (And (regular_i sig n (a::ls) i) (regular_i sig n rs i)) - (And (regular_i sig n ls i) (regular_i sig n (a::rs) i)). - -axiom regular_tail: ∀sig,n,l,i. - regular_i sig n l i → regular_i sig n (tail ? l) i. - -axiom regular_extend: ∀sig,n,l,i. - regular_i sig n l i → regular_i sig n (l@[all_blank sig n]) i. - -axiom all_blank_after_blank: ∀sig,n,l1,b,l2,i. - nth i ? (vec … b) (blank ?) = blank ? → - regular_i sig n (l1@b::l2) i → all_blanks_in ? (trace sig n i l2). - -lemma regular_trace_extl: ∀sig,n,a,ls,rs,i. - regular_trace sig n a ls rs i → - regular_trace sig n a (ls@[all_blank sig n]) rs i. -#sig #n #a #ls #rs #i * - [* #H1 #H2 % % // @(regular_extend … H1) - |* #H1 #H2 %2 % // @(regular_extend … H1) - ] -qed. - -lemma regular_cons_hd_rs: ∀sig,n.∀a:multi_sig sig n.∀ls,rs1,rs2,i. - regular_trace sig n a ls (rs1@rs2) i → - regular_trace sig n a ls (rs1@((hd ? rs2 (all_blank …))::(tail ? rs2))) i. -#sig #n #a #ls #rs1 #rs2 #i cases rs2 [2: #b #tl #H @H] -*[* #H1 >append_nil #H2 %1 % - [@H1 | whd in match (hd ???); @(regular_extend … rs1) //] - |* #H1 >append_nil #H2 %2 % - [@H1 | whd in match (hd ???); @(regular_extend … (a::rs1)) //] - ] -qed. - -lemma eq_trace_to_regular : ∀sig,n.∀a1,a2:multi_sig sig n.∀ls1,ls2,rs1,rs2,i. - nth i ? (vec … a1) (blank ?) = nth i ? (vec … a2) (blank ?) → - trace sig n i ls1 = trace sig n i ls2 → - trace sig n i rs1 = trace sig n i rs2 → - regular_trace sig n a1 ls1 rs1 i → - regular_trace sig n a2 ls2 rs2 i. -#sig #n #a1 #a2 #ls1 #ls2 #rs1 #rs2 #i #H1 #H2 #H3 #H4 -whd in match (regular_trace ??????); whd in match (regular_i ????); -whd in match (regular_i ?? rs2 ?); whd in match (regular_i ?? ls2 ?); -whd in match (regular_i ?? (a2::rs2) ?); whd in match (trace ????); -(Ht1a Hcur) in Ht2a; /2 by / - |#ls #a #rs #Htin #Hreg #Hreg2 -Ht1a cases (Ht1b … Htin) - [* #Hnb #Ht1 -Ht1b -Ht2a >Ht1 in Ht2b; >Htin #H - %{a} %{[ ]} %{ls} - %[%[%[%[%[@Hreg|@Hreg2]|@Hnb]|//]|//]|@H normalize % #H1 destruct (H1)] - |* - [(* we find the blank *) - * #ls1 * #b * #ls2 * * * #H1 #H2 #H3 #Ht1 - >Ht1 in Ht2b; #Hout -Ht1b - %{b} %{(a::ls1)} %{ls2} - %[%[%[%[%[>H1 in Hreg; #H @H - |#j #jneqi whd in match (hd ???); whd in match (tail ??); -

H1 //] - |@Hout normalize % normalize #H destruct (H) - ] - |* #b * #lss * * #H1 #H2 #Ht1 -Ht1b >Ht1 in Ht2a; - whd in match (left ??); whd in match (right ??); #Hout - %{(all_blank …)} %{(lss@[b])} %{[]} - %[%[%[%[%[

blank_all_blank //] - |

reverse_append >reverse_single @Hout normalize // - ] - ] - ] -qed. - -(******************************************************************************) - -definition shift_i_L ≝ λsig,n,i. - ncombf_r (multi_sig …) (shift_i sig n i) (all_blank sig n) · - mti sig n i · - extend ? (all_blank sig n). - -definition R_shift_i_L ≝ λsig,n,i,t1,t2. - (∀a,ls,rs. - t1 = midtape ? ls a rs → - ((∃rs1,b,rs2,a1,rss. - rs = rs1@b::rs2 ∧ - nth i ? (vec … b) (blank ?) = (blank ?) ∧ - (∀x. mem ? x rs1 → nth i ? (vec … x) (blank ?) ≠ (blank ?)) ∧ - shift_l sig n i (a::rs1) (a1::rss) ∧ - t2 = midtape (multi_sig sig n) ((reverse ? (a1::rss))@ls) b rs2) ∨ - (∃b,rss. - (∀x. mem ? x rs → nth i ? (vec … x) (blank ?) ≠ (blank ?)) ∧ - shift_l sig n i (a::rs) (rss@[b]) ∧ - t2 = midtape (multi_sig sig n) - ((reverse ? (rss@[b]))@ls) (all_blank sig n) [ ]))). - -definition R_shift_i_L_new ≝ λsig,n,i,t1,t2. - (∀a,ls,rs. - t1 = midtape ? ls a rs → - ∃rs1,b,rs2,rss. - b = hd ? rs2 (all_blank sig n) ∧ - nth i ? (vec … b) (blank ?) = (blank ?) ∧ - rs = rs1@rs2 ∧ - (∀x. mem ? x rs1 → nth i ? (vec … x) (blank ?) ≠ (blank ?)) ∧ - shift_l sig n i (a::rs1) rss ∧ - t2 = midtape (multi_sig sig n) ((reverse ? rss)@ls) b (tail ? rs2)). - -theorem sem_shift_i_L: ∀sig,n,i. shift_i_L sig n i ⊨ R_shift_i_L sig n i. -#sig #n #i -@(sem_seq_app ?????? - (sem_ncombf_r (multi_sig sig n) (shift_i sig n i)(all_blank sig n)) - (sem_seq ????? (ssem_mti sig n i) - (sem_extend ? (all_blank sig n)))) -#tin #tout * #t1 * * #Ht1a #Ht1b * #t2 * * #Ht2a #Ht2b * #Htout1 #Htout2 -#a #ls #rs cases rs - [#Htin %2 %{(shift_i sig n i a (all_blank sig n))} %{[ ]} - %[%[#x @False_ind | @daemon] - |lapply (Ht1a … Htin) -Ht1a -Ht1b #Ht1 - lapply (Ht2a … Ht1) -Ht2a -Ht2b #Ht2 >Ht2 in Htout1; - >Ht1 whd in match (left ??); whd in match (right ??); #Htout @Htout // - ] - |#a1 #rs1 #Htin - lapply (Ht1b … Htin) -Ht1a -Ht1b #Ht1 - lapply (Ht2b … Ht1) -Ht2a -Ht2b * - [(* a1 is blank *) * #H1 #H2 %1 - %{[ ]} %{a1} %{rs1} %{(shift_i sig n i a a1)} %{[ ]} - %[%[%[%[// |//] |#x @False_ind] | @daemon] - |>Htout2 [>H2 >reverse_single @Ht1 |>H2 >Ht1 normalize % #H destruct (H)] - ] - |* - [* #rs10 * #b * #rs2 * #rss * * * * #H1 #H2 #H3 #H4 - #Ht2 %1 - %{(a1::rs10)} %{b} %{rs2} %{(shift_i sig n i a a1)} %{rss} - %[%[%[%[>H1 //|//] |@H3] |@daemon ] - |>reverse_cons >associative_append - >H2 in Htout2; #Htout >Htout [@Ht2| >Ht2 normalize % #H destruct (H)] - ] - |* #b * #rss * * #H1 #H2 - #Ht2 %2 - %{(shift_i sig n i b (all_blank sig n))} %{(shift_i sig n i a a1::rss)} - %[%[@H1 |@daemon ] - |>Ht2 in Htout1; #Htout >Htout // - whd in match (left ??); whd in match (right ??); - >reverse_append >reverse_single >associative_append >reverse_cons - >associative_append // - ] - ] - ] - ] -qed. - -theorem sem_shift_i_L_new: ∀sig,n,i. - shift_i_L sig n i ⊨ R_shift_i_L_new sig n i. -#sig #n #i -@(Realize_to_Realize … (sem_shift_i_L sig n i)) -#t1 #t2 #H #a #ls #rs #Ht1 lapply (H a ls rs Ht1) * - [* #rs1 * #b * #rs2 * #a1 * #rss * * * * #H1 #H2 #H3 #H4 #Ht2 - %{rs1} %{b} %{(b::rs2)} %{(a1::rss)} - %[%[%[%[%[//|@H2]|@H1]|@H3]|@H4] | whd in match (tail ??); @Ht2] - |* #b * #rss * * #H1 #H2 #Ht2 - %{rs} %{(all_blank sig n)} %{[]} %{(rss@[b])} - %[%[%[%[%[//|@blank_all_blank]|//]|@H1]|@H2] | whd in match (tail ??); @Ht2] - ] -qed. - - -(******************************************************************************* -The following machine implements a full move of for a trace: we reach the left -border, shift the i-th trace and come back to the head position. *) - -(* this exclude the possibility that traces do not overlap: the head must -remain inside all traces *) - -definition mtiL ≝ λsig,n,i. - move_to_blank_L sig n i · - shift_i_L sig n i · - move_until ? L (no_head sig n). - -definition Rmtil ≝ λsig,n,i,t1,t2. - ∀ls,a,rs. - t1 = midtape (multi_sig sig n) ls a rs → - nth n ? (vec … a) (blank ?) = head ? → - (∀i.regular_trace sig n a ls rs i) → - (* next: we cannot be on rightof on trace i *) - (nth i ? (vec … a) (blank ?) = (blank ?) - → nth i ? (vec … (hd ? rs (all_blank …))) (blank ?) ≠ (blank ?)) → - no_head_in … ls → - no_head_in … rs → - (∃ls1,a1,rs1. - t2 = midtape (multi_sig …) ls1 a1 rs1 ∧ - (∀i.regular_trace … a1 ls1 rs1 i) ∧ - (∀j. j ≤ n → j ≠ i → to_blank_i ? n j (a1::ls1) = to_blank_i ? n j (a::ls)) ∧ - (∀j. j ≤ n → j ≠ i → to_blank_i ? n j rs1 = to_blank_i ? n j rs) ∧ - (to_blank_i ? n i ls1 = to_blank_i ? n i (a::ls)) ∧ - (to_blank_i ? n i (a1::rs1)) = to_blank_i ? n i rs). - -theorem sem_Rmtil: ∀sig,n,i. i < n → mtiL sig n i ⊨ Rmtil sig n i. -#sig #n #i #lt_in -@(sem_seq_app ?????? - (sem_move_to_blank_L … ) - (sem_seq ????? (sem_shift_i_L_new …) - (ssem_move_until_L ? (no_head sig n)))) -#tin #tout * #t1 * * #_ #Ht1 * #t2 * #Ht2 * #_ #Htout -(* we start looking into Rmitl *) -#ls #a #rs #Htin (* tin is a midtape *) -#Hhead #Hreg #no_rightof #Hnohead_ls #Hnohead_rs -cut (regular_i sig n (a::ls) i) - [cases (Hreg i) * // - cases (true_or_false (nth i ? (vec … a) (blank ?) == (blank ?))) #Htest - [#_ @daemon (* absurd, since hd rs non e' blank *) - |#H #_ @daemon]] #Hreg1 -lapply (Ht1 … Htin Hreg1 ?) [#j #_ @Hreg] -Ht1 -Htin -* #b * #ls1 * #ls2 * * * * * #reg_ls1_i #reg_ls1_j #Hno_blankb #Hhead #Hls1 #Ht1 -lapply (Ht2 … Ht1) -Ht2 -Ht1 -* #rs1 * #b0 * #rs2 * #rss * * * * * #Hb0 #Hb0blank #Hrs1 #Hrs1b #Hrss #Ht2 -(* we need to recover the position of the head of the emulated machine - that is the head of ls1. This is somewhere inside rs1 *) -cut (∃rs11. rs1 = (reverse ? ls1)@rs11) - [cut (ls1 = [ ] ∨ ∃aa,tlls1. ls1 = aa::tlls1) - [cases ls1 [%1 // | #aa #tlls1 %2 %{aa} %{tlls1} //]] * - [#H1ls1 %{rs1} >H1ls1 // - |* #aa * #tlls1 #H1ls1 >H1ls1 in Hrs1; - cut (aa = a) [>H1ls1 in Hls1; #H @(to_blank_hd … H)] #eqaa >eqaa - #Hrs1_aux cases (compare_append … (sym_eq … Hrs1_aux)) #l * - [* #H1 #H2 %{l} @H1 - |(* this is absurd : if l is empty, the case is as before. - if l is not empty then it must start with a blank, since it is the - first character in rs2. But in this case we would have a blank - inside ls1=a::tls1 that is absurd *) - @daemon - ]]] - * #rs11 #H1 -cut (rs = rs11@rs2) - [@(injective_append_l … (reverse … ls1)) >Hrs1 H2 >trace_append @mem_append_l2 - lapply Hb0 cases rs2 - [whd in match (hd ???); #H >H in H3; whd in match (no_head ???); - >all_blank_n normalize -H #H destruct (H); @False_ind - |#c #r #H4 %1 >H4 // - ] - |* - [(* we reach the head position *) - (* cut (trace sig n j (a1::ls20)=trace sig n j (ls1@b::ls2)) *) - * #ls10 * #a1 * #ls20 * * * #Hls20 #Ha1 #Hnh #Htout - cut (∀j.j ≠ i → - trace sig n j (reverse (multi_sig sig n) rs1@b::ls2) = - trace sig n j (ls10@a1::ls20)) - [#j #ineqj >append_cons trace_def reverse_map >map_append @eq_f @Hls20 ] - #Htracej - cut (trace sig n i (reverse (multi_sig sig n) (rs1@[b0])@ls2) = - trace sig n i (ls10@a1::ls20)) - [>trace_def Hcut - lapply (trace_shift … lt_in … Hrss) [//] whd in match (tail ??); #Htr reverse_map >map_append trace_def in ⊢ (%→?); trace_def in ⊢ (%→?); H1 in Htracei; >reverse_append >reverse_single >reverse_append - >reverse_reverse >associative_append >associative_append - @daemon - ] #H3 - cut (∀j. j ≠ i → - trace sig n j (reverse (multi_sig sig n) ls10@rs2) = trace sig n j rs) - [#j #jneqi @(injective_append_l … (trace sig n j (reverse ? ls1))) - >map_append >map_append >Hrs1 >H1 >associative_append - eqji (* by cases wether a1 is blank *) - @daemon - |#jneqi lapply (reg_ls1_j … jneqi) #H4 - >reverse_cons >associative_append >Hb0 @regular_cons_hd_rs - @(eq_trace_to_regular … H4) - [(proj2 … (H2 … jneqi)) >hd_trace % - |(proj2 … (H2 … jneqi)) >tail_trace % - |@sym_eq @Hrs_j // - ] - ]] - |#j #lejn #jneqi <(Hls1 … lejn) - >to_blank_i_def >to_blank_i_def @eq_f @sym_eq @(proj2 … (H2 j jneqi))] - |#j #lejn #jneqi >reverse_cons >associative_append >Hb0 - to_blank_i_def >to_blank_i_def @eq_f @Hrs_j //] - |<(Hls1 i) [2:@lt_to_le //] - lapply (all_blank_after_blank … reg_ls1_i) - [@(\P ?) @daemon] #allb_ls2 - whd in match (to_blank_i ????); <(proj2 … H3) - @daemon ] - |>reverse_cons >associative_append - cut (to_blank_i sig n i rs = to_blank_i sig n i (rs11@[b0])) [@daemon] - #Hcut >Hcut >(to_blank_i_chop … b0 (a1::reverse …ls10)) [2: @Hb0blank] - >to_blank_i_def >to_blank_i_def @eq_f - >trace_def >trace_def @injective_reverse >reverse_map >reverse_cons - >reverse_reverse >reverse_map >reverse_append >reverse_single @sym_eq - @(proj1 … H3) - ] - |(*we do not find the head: this is absurd *) - * #b1 * #lss * * #H2 @False_ind - cut (∀x0. mem ? x0 (trace sig n n (b0::reverse ? rss@ls2)) → x0 ≠ head ?) - [@daemon] -H2 #H2 - lapply (trace_shift_neq sig n i n … lt_in … Hrss) - [@lt_to_not_eq @lt_in | // ] - #H3 @(absurd - (nth n ? (vec … (hd ? (ls1@[b]) (all_blank sig n))) (blank ?) = head ?)) - [>Hhead // - |@H2 >trace_def %2 H3 >H1 >trace_def >reverse_map >reverse_cons >reverse_append - >reverse_reverse >associative_append