From ebf0b4572d0d3252dbbc7b769df676b76bd38c22 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Sun, 26 Jan 2014 20:12:59 +0000 Subject: [PATCH] nat: we added a non-indexed theorem lambda: the corrispondence between terms by level and terms by depth is proved! --- matita/matita/lib/arithmetics/nat.ma | 285 +++++++++--------- .../lib/lambda/levels/interpretations.ma | 40 ++- .../lib/lambda/levels/iterated_abstraction.ma | 32 ++ .../lib/lambda/terms/iterated_abstraction.ma | 2 +- 4 files changed, 205 insertions(+), 154 deletions(-) create mode 100644 matita/matita/lib/lambda/levels/iterated_abstraction.ma diff --git a/matita/matita/lib/arithmetics/nat.ma b/matita/matita/lib/arithmetics/nat.ma index 3a69562d6..cca478280 100644 --- a/matita/matita/lib/arithmetics/nat.ma +++ b/matita/matita/lib/arithmetics/nat.ma @@ -86,7 +86,7 @@ interpretation "natural minus" 'minus x y = (minus x y). theorem nat_case: ∀n:nat.∀P:nat → Prop. (n=O → P O) → (∀m:nat. n= S m → P (S m)) → P n. -#n #P elim n /2/ qed. +#n #P (elim n) /2/ qed. theorem nat_elim2 : ∀R:nat → nat → Prop. @@ -94,7 +94,7 @@ theorem nat_elim2 : → (∀n:nat. R (S n) O) → (∀n,m:nat. R n m → R (S n) (S m)) → ∀n,m:nat. R n m. -#R #ROn #RSO #RSS #n elim n // #n0 #Rn0m #m cases m /2/ qed. +#R #ROn #RSO #RSS #n (elim n) // #n0 #Rn0m #m (cases m) /2/ qed. lemma le_gen: ∀P:nat → Prop.∀n.(∀i. i ≤ n → P i) → P n. /2/ qed. @@ -115,25 +115,25 @@ theorem plus_O_n: ∀n:nat. n = 0 + n. // qed. theorem plus_n_O: ∀n:nat. n = n + 0. -#n elim n normalize // qed. +#n (elim n) normalize // qed. theorem plus_n_Sm : ∀n,m:nat. S (n+m) = n + S m. -#n elim n normalize // qed. +#n (elim n) normalize // qed. theorem commutative_plus: commutative ? plus. -#n elim n normalize // qed. +#n (elim n) normalize // qed. theorem associative_plus : associative nat plus. -#n elim n normalize // qed. +#n (elim n) normalize // qed. theorem assoc_plus1: ∀a,b,c. c + (b + a) = b + c + a. // qed. theorem injective_plus_r: ∀n:nat.injective nat nat (λm.n+m). -#n elim n normalize /3 width=1 by injective_S/ qed. +#n (elim n) normalize /3/ qed. theorem injective_plus_l: ∀n:nat.injective nat nat (λm.m+n). -/2 width=2 by injective_plus_r/ qed. +/2/ qed. theorem times_Sn_m: ∀n,m:nat. m+n*m = S n*m. // qed. @@ -142,23 +142,23 @@ theorem times_O_n: ∀n:nat. 0 = 0 * n. // qed. theorem times_n_O: ∀n:nat. 0 = n * 0. -#n elim n // qed. +#n (elim n) // qed. theorem times_n_Sm : ∀n,m:nat. n+(n*m) = n*(S m). -#n elim n normalize // qed. +#n (elim n) normalize // qed. theorem commutative_times : commutative nat times. -#n elim n normalize // qed. +#n (elim n) normalize // qed. theorem distributive_times_plus : distributive nat times plus. -#n elim n normalize // qed. +#n (elim n) normalize // qed. theorem distributive_times_plus_r : ∀a,b,c:nat. (b+c)*a = b*a + c*a. // qed. theorem associative_times: associative nat times. -#n elim n normalize // qed. +#n (elim n) normalize // qed. lemma times_times: ∀x,y,z. x*(y*z) = y*(x*z). // qed. @@ -170,16 +170,16 @@ theorem minus_S_S: ∀n,m:nat.S n - S m = n -m. // qed. theorem minus_O_n: ∀n:nat.0=0-n. -#n cases n // qed. +#n (cases n) // qed. theorem minus_n_O: ∀n:nat.n=n-0. -#n cases n // qed. +#n (cases n) // qed. theorem minus_n_n: ∀n:nat.0=n-n. -#n elim n // qed. +#n (elim n) // qed. theorem minus_Sn_n: ∀n:nat. S 0 = (S n)-n. -#n elim n normalize // qed. +#n (elim n) normalize // qed. theorem eq_minus_S_pred: ∀n,m. n - (S m) = pred(n -m). @nat_elim2 normalize // qed. @@ -189,7 +189,7 @@ lemma plus_plus_comm_23: ∀x,y,z. x + y + z = x + z + y. lemma discr_plus_xy_minus_xz: ∀x,z,y. x + y = x - z → y = 0. #x elim x -x // #x #IHx * normalize -[ #y #H @(IHx 0) plus_n_Sm #H lapply (IHx … H) -x -z #H destruct ] qed-. @@ -197,69 +197,68 @@ qed-. (* Negated equalities *******************************************************) theorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m. -/2 width=3 by not_to_not/ qed. +/2/ qed. theorem not_eq_O_S : ∀n:nat. 0 ≠ S n. -#n @nmk #eqOS change with (not_zero 0) >eqOS // qed. +#n @nmk #eqOS (change with (not_zero O)) >eqOS // qed. theorem not_eq_n_Sn: ∀n:nat. n ≠ S n. -#n elim n /2 width=1 by not_eq_S/ qed. +#n (elim n) /2/ qed. (* Atomic conclusion *******************************************************) (* not_zero *) theorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m. -#n #m #Hlt elim Hlt // qed. +#n #m #Hlt (elim Hlt) // qed. (* le *) theorem le_S_S: ∀n,m:nat. n ≤ m → S n ≤ S m. -#n #m #lenm elim lenm /2 width=1 by le_S/ qed. +#n #m #lenm (elim lenm) /2/ qed. theorem le_O_n : ∀n:nat. 0 ≤ n. -#n elim n /2 width=1 by le_S/ qed. +#n (elim n) /2/ qed. theorem le_n_Sn : ∀n:nat. n ≤ S n. -/2 width=1 by le_S/ qed. +/2/ qed. theorem transitive_le : transitive nat le. -#a #b #c #leab #lebc elim lebc /2 width=1 by le_S/ +#a #b #c #leab #lebc (elim lebc) /2/ qed. theorem le_pred_n : ∀n:nat. pred n ≤ n. -#n elim n // qed. +#n (elim n) // qed. theorem monotonic_pred: monotonic ? le pred. -#n #m #lenm elim lenm /2 width=3 by transitive_le/ qed. +#n #m #lenm (elim lenm) /2/ qed. theorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m. (* demo *) -/2 width=1 by monotonic_pred/ qed-. +/2/ qed-. theorem monotonic_le_plus_r: ∀n:nat.monotonic nat le (λm.n + m). -#n #a #b elim n normalize // -#m #H #leab /3 width=1 by le_S_S/ -qed. +#n #a #b (elim n) normalize // +#m #H #leab @le_S_S /2/ qed. theorem monotonic_le_plus_l: ∀m:nat.monotonic nat le (λn.n + m). -/2 width=1 by monotonic_le_plus_r/ qed. +/2/ qed. theorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2 → m1 ≤ m2 → n1 + m1 ≤ n2 + m2. #n1 #n2 #m1 #m2 #len #lem @(transitive_le ? (n1+m2)) -/2 width=1 by monotonic_le_plus_l, monotonic_le_plus_r/ qed. +/2/ qed. theorem le_plus_n :∀n,m:nat. m ≤ n + m. -/2 width=1 by monotonic_le_plus_l/ qed. +/2/ qed. lemma le_plus_a: ∀a,n,m. n ≤ m → n ≤ a + m. -/2 width=1 by le_plus/ qed. +/2/ qed. lemma le_plus_b: ∀b,n,m. n + b ≤ m → n ≤ m. -/2 width=3 by transitive_le/ qed. +/2/ qed. theorem le_plus_n_r :∀n,m:nat. m ≤ m + n. /2/ qed. @@ -268,46 +267,47 @@ theorem eq_plus_to_le: ∀n,m,p:nat.n=m+p → m ≤ n. // qed-. theorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m. -#a elim a normalize /3 width=1 by monotonic_pred/ qed. +#a (elim a) normalize /3/ qed. theorem le_plus_to_le_r: ∀a,n,m. n + a ≤ m +a → n ≤ m. -/2 width=2 by le_plus_to_le/ qed-. +/2/ qed-. theorem monotonic_le_times_r: ∀n:nat.monotonic nat le (λm. n * m). -#n #x #y #lexy elim n normalize /2 width=1 by le_plus/ +#n #x #y #lexy (elim n) normalize//(* lento /2/*) +#a #lea @le_plus // qed. theorem le_times: ∀n1,n2,m1,m2:nat. n1 ≤ n2 → m1 ≤ m2 → n1*m1 ≤ n2*m2. -#n1 #n2 #m1 #m2 #len #lem @(transitive_le ? (n1*m2)) -/2 width=1 by monotonic_le_times_r/ +#n1 #n2 #m1 #m2 #len #lem @(transitive_le ? (n1*m2)) /2/ qed. (* interessante *) theorem lt_times_n: ∀n,m:nat. O < n → m ≤ n*m. -/2 width=1 by monotonic_le_times_r/ qed. +#n #m #H /2/ qed. theorem le_times_to_le: ∀a,n,m. O < a → a * n ≤ a * m → n ≤ m. #a @nat_elim2 normalize [// |#n #H1 #H2 - @(transitive_le ? (a*S n)) /2 width=1 by monotonic_le_times_r/ - |#n #m #H #lta #le /4 width=2 by le_plus_to_le, le_S_S/ + @(transitive_le ? (a*S n)) /2/ + |#n #m #H #lta #le + @le_S_S @H /2/ ] qed-. theorem le_plus_minus_m_m: ∀n,m:nat. n ≤ (n-m)+m. -#n elim n // #a #Hind #m cases m // normalize #n /2 width=1 by le_S_S/ +#n (elim n) // #a #Hind #m (cases m) // normalize #n/2/ qed. theorem le_plus_to_minus_r: ∀a,b,c. a + b ≤ c → a ≤ c -b. -#a #b #c #H @(le_plus_to_le_r … b) /2 width=3 by transitive_le/ +#a #b #c #H @(le_plus_to_le_r … b) /2/ qed. lemma lt_to_le: ∀x,y. x < y → x ≤ y. -/2 width=2 by le_plus_b/ qed. +/2 width=2/ qed. lemma inv_eq_minus_O: ∀x,y. x - y = 0 → x ≤ y. // qed-. @@ -319,84 +319,84 @@ qed. (* lt *) theorem transitive_lt: transitive nat lt. -#a #b #c #ltab #ltbc elim ltbc /2 width=1 by le_S/ +#a #b #c #ltab #ltbc (elim ltbc) /2/ qed. theorem lt_to_le_to_lt: ∀n,m,p:nat. n < m → m ≤ p → n < p. -#n #m #p #H #H1 elim H1 /2 width=3 by transitive_lt/ qed. +#n #m #p #H #H1 (elim H1) /2/ qed. theorem le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p. -#n #m #p #H elim H /3 width=3 by transitive_lt/ qed. +#n #m #p #H (elim H) /3/ qed. theorem lt_S_to_lt: ∀n,m. S n < m → n < m. -/2 width=3 by transitive_lt/ qed. +/2/ qed. theorem ltn_to_ltO: ∀n,m:nat. n < m → 0 < m. -/2 width=3 by lt_to_le_to_lt/ qed. +/2/ qed. theorem lt_O_S : ∀n:nat. O < S n. -/2 width=1 by ltn_to_ltO/ qed. +/2/ qed. theorem monotonic_lt_plus_r: ∀n:nat.monotonic nat lt (λm.n+m). -/2 width=1 by monotonic_le_plus_r/ qed. +/2/ qed. theorem monotonic_lt_plus_l: ∀n:nat.monotonic nat lt (λm.m+n). -/2 width=1 by monotonic_le_plus_r/ qed. +/2/ qed. theorem lt_plus: ∀n,m,p,q:nat. n < m → p < q → n + p < m + q. #n #m #p #q #ltnm #ltpq -@(transitive_lt ? (n+q)) /2 width=1 by monotonic_lt_plus_l, monotonic_le_plus_r/ qed. +@(transitive_lt ? (n+q))/2/ qed. theorem lt_plus_to_lt_l :∀n,p,q:nat. p+n < q+n → p(plus_minus_m_m … Hlecb) /2 width=1 by monotonic_le_plus_l/ +#a #b #c #Hlecb #H >(plus_minus_m_m … Hlecb) /2/ qed. theorem le_plus_to_minus: ∀n,m,p. n ≤ p+m → n-m ≤ p. -/2 width=1 by monotonic_le_minus_l/ qed. +#n #m #p #lep /2/ qed. theorem monotonic_le_minus_r: ∀p,q,n:nat. q ≤ p → n-p ≤ n-q. #p #q #n #lepq @le_plus_to_minus -@(transitive_le … (le_plus_minus_m_m ? q)) /2 width=1 by monotonic_le_plus_r/ +@(transitive_le … (le_plus_minus_m_m ? q)) /2/ qed. theorem increasing_to_le: ∀f:nat → nat. increasing f → ∀m:nat.∃i.m ≤ f i. -#f #incr #m elim m /2 width=2 by ex_intro/ -#n * #a #lenfa -@(ex_intro ?? (S a)) /2 width=3 by le_to_lt_to_lt/ +#f #incr #m (elim m) /2/#n * #a #lenfa +@(ex_intro ?? (S a)) /2/ qed. +(* thus is le_plus +lemma le_plus_compatible: ∀x1,x2,y1,y2. x1 ≤ y1 → x2 ≤ y2 → x1 + x2 ≤ y1 + y2. +#x1 #y1 #x2 #y2 #H1 #H2 /2/ @le_plus // /2/ /3 by le_minus_to_plus, monotonic_le_plus_r, transitive_le/ qed. +*) + lemma minus_le: ∀x,y. x - y ≤ x. -/2 width=1 by monotonic_le_minus_l/ qed. +/2 width=1/ qed. (* lt *) theorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → neq_minus_O [ >eq_minus_O // @monotonic_le_times_r ] - /2 width=1 by lt_to_le/ + [> eq_minus_O [2:/2/] >eq_minus_O // + @monotonic_le_times_r /2/ |@sym_eq (applyS plus_to_minus) eq_minus_O /2 width=1 by eq_minus_O, monotonic_le_minus_l/ (* >eq_minus_O // *) + #H >eq_minus_O /2/ (* >eq_minus_O // *) ] qed. @@ -753,17 +749,17 @@ theorem minus_minus: ∀n,m,p:nat. p ≤ m → m ≤ n → qed. lemma minus_minus_associative: ∀x,y,z. z ≤ y → y ≤ x → x - (y - z) = x - y + z. -/2 width=1 by minus_minus/ qed. +/2 width=1 by minus_minus/ qed-. lemma minus_minus_comm: ∀a,b,c. a - b - c = a - c - b. -/3 width=1 by monotonic_le_minus_l, le_to_le_to_eq/ qed. +/3 by monotonic_le_minus_l, le_to_le_to_eq/ qed. lemma minus_le_minus_minus_comm: ∀b,c,a. c ≤ b → a - (b - c) = a + c - b. #b #c #a #H >(plus_minus_m_m b c) in ⊢ (? ? ?%); // qed. lemma minus_minus_m_m: ∀m,n. n ≤ m → m - (m - n) = n. -/2 width=1 by minus_le_minus_minus_comm/ qed. +/2 width=1/ qed. lemma minus_plus_plus_l: ∀x,y,h. (x + h) - (y + h) = x - y. // qed. @@ -771,6 +767,9 @@ lemma minus_plus_plus_l: ∀x,y,h. (x + h) - (y + h) = x - y. lemma plus_minus_plus_plus_l: ∀z,x,y,h. z + (x + h) - (y + h) = z + x - y. // qed. +lemma minus_plus_minus_l: ∀x,y,z. y ≤ z → (z + x) - (z - y) = x + y. +/2 width=1 by minus_minus_associative/ qed-. + (* Stilll more atomic conclusion ********************************************) (* le *) diff --git a/matita/matita/lib/lambda/levels/interpretations.ma b/matita/matita/lib/lambda/levels/interpretations.ma index 0c06ab442..99643b17a 100644 --- a/matita/matita/lib/lambda/levels/interpretations.ma +++ b/matita/matita/lib/lambda/levels/interpretations.ma @@ -17,7 +17,7 @@ include "lambda/notation/functions/forward_3.ma". include "lambda/notation/functions/backward_1.ma". include "lambda/notation/functions/backward_3.ma". include "lambda/terms/iterated_abstraction.ma". -include "lambda/levels/term.ma". +include "lambda/levels/iterated_abstraction.ma". (* INTERPRETATIONS **********************************************************) @@ -33,6 +33,10 @@ interpretation "forward interpretation (term by depth) general" interpretation "forward interpretation (term by depth)" 'Forward M = (bylevel O O M). +lemma bylevel_abst: ∀i,h,d,M. ⇑[d, h] 𝛌i. M = ⇑[i+d, i+h] M. +#i elim i -i normalize // +qed. + let rec bydepth h d M on M ≝ match M with [ LVRef i e ⇒ 𝛌i.#(tri … e (d+i-h) (d+i-h-e-1) e e) | LAppl i C A ⇒ 𝛌i.@(bydepth h (d+i) C).(bydepth h (d+i) A) @@ -44,27 +48,43 @@ interpretation "backward interpretation (term by level) general" interpretation "backward interpretation (term by level)" 'Backward M = (bydepth O O M). -theorem by_depth_level_gen: ∀M,e,d,h. d ≤ e + h → ⇓[e, e+h-d] ⇑[d, h] M = 𝛌h.M. +lemma by_depth_level_gen: ∀M,e,d,h. d ≤ e + h → ⇓[e, e+h-d] ⇑[d, h] M = 𝛌h.M. #M elim M -M normalize [ #i #e #d #h #Hdeh >(minus_minus_m_m … Hdeh) elim (lt_or_eq_or_gt i d) #Hid [ >(tri_lt ???? … Hid) >(tri_lt ???? d (d-i-1)) - [ >minus_minus_associative /2 width=1 by monotonic_le_minus_r/ - minus_minus_associative /2 width=1 by lt_to_le/ - | /2 width=1 by monotonic_lt_minus_l/ - ] + /5 width=1 by minus_le_minus_minus_comm, monotonic_lt_minus_l, eq_f/ | destruct >(tri_eq ???? …) >(tri_eq ???? …) // | >(tri_gt ???? … Hid) >(tri_gt ???? … Hid) // ] | #A #IHA #e #d #h #Hdeh lapply (IHA e (d+1) (h+1) ?) -IHA /2 width=1 by le_S_S, eq_f2/ | #C #A #IHC #IHA #e #d #h #Hdeh - lapply (IHC (e+h) d 0 ?) -IHC - lapply (IHA (e+h) d 0 ?) -IHA + lapply (IHC (e+h) d 0 ?) -IHC // lapply (IHA (e+h) d 0 ?) -IHA // normalize /2 width=1 by/ ] -qed. +qed-. -lemma by_depth_level: ∀M. ⇓⇑M = M. +theorem by_depth_level: ∀M. ⇓⇑M = M. #M lapply (by_depth_level_gen M 0 0 0 ?) normalize // qed. + +lemma by_level_depth_gen: ∀M,e,d,h. d ≤ e → ⇑[d, h] ⇓[e, e-d] M = 𝛌h.M. +#M elim M -M +[ #i #k #e #d #h #Hde >bylevel_abst normalize >(minus_plus_minus_l … Hde) + elim (lt_or_eq_or_gt k (i+d)) #Hkid + [ >(tri_lt ???? … Hkid) >(tri_lt ???? (i+d) (i+d-k-1)) + /5 width=1 by minus_le_minus_minus_comm, monotonic_lt_minus_l, eq_f/ + | destruct >(tri_eq ???? …) >(tri_eq ???? …) // + | >(tri_gt ???? … Hkid) >(tri_gt ???? … Hkid) // + ] +| #i #C #A #IHC #IHA #e #d #h #Hdeh >bylevel_abst normalize + lapply (IHC (e+i) (i+d) 0 ?) -IHC /2 width=1 by monotonic_le_plus_r/ + lapply (IHA (e+i) (i+d) 0 ?) -IHA /2 width=1 by monotonic_le_plus_r/ + /3 width=1 by eq_f3, eq_f2/ +] +qed-. + +theorem by_level_depth: ∀M. ⇑⇓M = M. +#M lapply (by_level_depth_gen M 0 0 0 ?) // +qed. diff --git a/matita/matita/lib/lambda/levels/iterated_abstraction.ma b/matita/matita/lib/lambda/levels/iterated_abstraction.ma new file mode 100644 index 000000000..c2ad35be7 --- /dev/null +++ b/matita/matita/lib/lambda/levels/iterated_abstraction.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "lambda/levels/term.ma". + +(* ITERATED ABSTRACTION *****************************************************) + +definition labst: nat → lterm → lterm ≝ λh,M. match M with +[ LVRef i e ⇒ {i+h}§e +| LAppl i C A ⇒ {i+h}@C.A +]. + +interpretation "iterated abstraction (term by level)" + 'AnnotatedAbstraction h M = (labst h M). + +interpretation "abstraction (term by level)" + 'Abstraction M = (labst (S O) M). + +lemma labst_O: ∀M. 𝛌0.M = M. +* normalize // +qed. diff --git a/matita/matita/lib/lambda/terms/iterated_abstraction.ma b/matita/matita/lib/lambda/terms/iterated_abstraction.ma index 78d150ba3..a4aa4f46f 100644 --- a/matita/matita/lib/lambda/terms/iterated_abstraction.ma +++ b/matita/matita/lib/lambda/terms/iterated_abstraction.ma @@ -14,7 +14,7 @@ include "lambda/terms/term.ma". -(* GENERALIZED ABSTRACTION **************************************************) +(* ITERATED ABSTRACTION *****************************************************) let rec abst d M on d ≝ match d with [ O ⇒ M -- 2.39.2