]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda/ext.ma
more notation and all-purpose lemmas
[helm.git] / matita / matita / lib / lambda / ext.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "lambda/subst.ma".
16 include "basics/list.ma".
17 include "lambda/lambda_notation.ma".
18
19 (* MATTER CONCERNING STRONG NORMALIZATION TO BE PUT ELSEWHERE *****************)
20
21 (* arithmetics ****************************************************************)
22
23 lemma arith1: ∀x,y. (S x) ≰ (S y) → x ≰ y.
24 #x #y #HS @nmk (elim HS) -HS /3/
25 qed.
26
27 lemma arith2: ∀i,p,k. k ≤ i → i + p - (k + p) = i - k.
28 #i #p #k #H @plus_to_minus
29 >commutative_plus >(commutative_plus k) >associative_plus @eq_f /2/
30 qed.
31
32 lemma arith3: ∀x,y,z. x ≰ y → x + z ≰ y + z.
33 #x #y #z #H @nmk (elim H) -H /3/
34 qed.
35
36 lemma arith4: ∀x,y. S x ≰ y → x≠y → y < x.
37 #x #y #H1 #H2 lapply (not_le_to_lt … H1) -H1 #H1 @not_eq_to_le_to_lt /2/
38 qed.
39
40 lemma arith5: ∀x,y. x < y → S (y - 1) ≰ x.
41 #x #y #H @lt_to_not_le <minus_Sn_m /2/
42 qed.
43
44 (* lists **********************************************************************)
45
46 (* all(?,P,l) holds when P holds for all members of l *)
47 let rec all (A:Type[0]) (P:A→Prop) l on l ≝ match l with 
48    [ nil        ⇒ True
49    | cons hd tl ⇒ P hd ∧ all A P tl
50    ].
51
52 lemma all_hd: ∀A:Type[0]. ∀P:A→Prop. ∀a. P a → ∀l. all … P l → P (hd … l a).
53 #A #P #a #Ha #l elim l -l [ #_ @Ha | #b #l #_ #Hl elim Hl // ]
54 qed.
55
56 lemma all_tl: ∀A:Type[0]. ∀P:A→Prop. ∀l. all … P l →  all … P (tail … l).
57 #A #P #l elim l -l // #b #l #IH #Hl elim Hl //
58 qed.
59
60 lemma all_nth: ∀A:Type[0]. ∀P:A→Prop. ∀a. P a → ∀i,l. all … P l → P (nth i … l a).
61 #A #P #a #Ha #i elim i -i [ @all_hd // | #i #IH #l #Hl @IH /2/ ]
62 qed.
63
64 lemma all_append: ∀A,P,l2,l1. all A P l1 → all A P l2 → all A P (l1 @ l2).
65 #A #P #l2 #l1 (elim l1) -l1 (normalize) // #hd #tl #IH1 #H (elim H) /3/
66 qed.
67
68 (* all2(?,P,l1,l2) holds when P holds for all paired members of l1 and l2 *)
69 let rec all2 (A:Type[0]) (P:A→A→Prop) l1 l2 on l1 ≝ match l1 with
70    [ nil          ⇒ l2 = nil ?
71    | cons hd1 tl1 ⇒ match l2 with
72       [ nil          ⇒ False
73       | cons hd2 tl2 ⇒ P hd1 hd2 ∧ all2 A P tl1 tl2
74       ]
75    ].
76
77 lemma length_append: ∀A. ∀(l2,l1:list A). |l1@l2| = |l1| + |l2|.
78 #A #l2 #l1 (elim l1) -l1 (normalize) //
79 qed.
80
81 (* terms **********************************************************************)
82
83 (* Appl F l generalizes App applying F to a list of arguments
84  * The head of l is applied first
85  *)
86 let rec Appl F l on l ≝ match l with 
87    [ nil ⇒ F
88    | cons A D ⇒ Appl (App F A) D  
89    ].
90
91 lemma appl_append: ∀N,l,M. Appl M (l @ [N]) = App (Appl M l) N.
92 #N #l (elim l) -l // #hd #tl #IHl #M >IHl //
93 qed.
94
95 (* FG: not needed for now 
96 (* nautral terms *)
97 inductive neutral: T → Prop ≝
98    | neutral_sort: ∀n.neutral (Sort n)
99    | neutral_rel: ∀i.neutral (Rel i)
100    | neutral_app: ∀M,N.neutral (App M N)
101 .
102 *)
103
104 (* substitution ***************************************************************)
105
106 (* FG: do we need this? 
107 definition lift0 ≝ λp,k,M . lift M p k. (**) (* remove definition *)
108
109 lemma lift_appl: ∀p,k,l,F. lift (Appl F l) p k = 
110                              Appl (lift F p k) (map … (lift0 p k) l). 
111 #p #k #l (elim l) -l /2/ #A #D #IHl #F >IHl //
112 qed.
113 *)
114
115 lemma lift_rel_lt: ∀i,p,k. (S i) ≤ k → lift (Rel i) k p = Rel i.
116 #i #p #k #Hik normalize >(le_to_leb_true … Hik) //
117 qed.
118
119 lemma lift_rel_ge: ∀i,p,k. (S i) ≰ k → lift (Rel i) k p = Rel (i+p).
120 #i #p #k #Hik normalize >(lt_to_leb_false (S i) k) /2/
121 qed.
122
123 lemma lift_app: ∀M,N,k,p.
124                 lift (App M N) k p = App (lift M k p) (lift N k p).
125 // qed.
126
127 lemma lift_lambda: ∀N,M,k,p. lift (Lambda N M) k p = 
128                    Lambda (lift N k p) (lift M (k + 1) p).
129 // qed.
130
131 lemma lift_prod: ∀N,M,k,p.
132                  lift (Prod N M) k p = Prod (lift N k p) (lift M (k + 1) p).
133 // qed.
134
135 lemma subst_app: ∀M,N,k,L. (App M N)[k≝L] = App M[k≝L] N[k≝L].
136 // qed.
137
138 lemma subst_lambda: ∀N,M,k,L. (Lambda N M)[k≝L] = Lambda N[k≝L] M[k+1≝L].
139 // qed.
140
141 lemma subst_prod: ∀N,M,k,L. (Prod N M)[k≝L] = Prod N[k≝L] M[k+1≝L].
142 // qed.
143
144
145 axiom lift_subst_lt: ∀A,B,i,j,k. lift (B[j≝A]) (j+k) i =
146                      (lift B (j+k+1) i)[j≝lift A k i].
147
148 (* telescopic delifting substitution of l in M.
149  * Rel 0 is replaced with the head of l
150  *)
151 let rec tsubst M l on l ≝ match l with
152    [ nil      ⇒ M
153    | cons A D ⇒ (tsubst M[0≝A] D)
154    ]. 
155
156 interpretation "telescopic substitution" 'Subst1 M l = (tsubst M l).
157
158 lemma tsubst_refl: ∀l,t. (lift t 0 (|l|))[l] = t.
159 #l (elim l) -l (normalize) // #hd #tl #IHl #t cut (S (|tl|) = |tl| + 1) // (**) (* eliminate cut *)
160 qed.
161
162 lemma tsubst_sort: ∀n,l. (Sort n)[l] = Sort n.
163 //
164 qed.