]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground/relocation/tr_pat.ma
update in ground
[helm.git] / matita / matita / contribs / lambdadelta / ground / relocation / tr_pat.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 "ground/arith/pnat_le_plus.ma".
16 include "ground/relocation/pr_pat_lt.ma".
17 include "ground/relocation/tr_map.ma".
18
19 (* TOTAL RELOCATION MAPS ****************************************************)
20
21 (* Constructions with pr_pat ***********************************************)
22
23 (*** at_O1 *)
24 lemma tr_pat_unit_sn: ∀i2,f. @⧣❨𝟏,𝐭❨i2⨮f❩❩ ≘ i2.
25 #i2 elim i2 -i2 /2 width=5 by pr_pat_refl, pr_pat_next/
26 qed.
27
28 (*** at_S1 *)
29 lemma tr_pat_succ_sn: ∀p,f,i1,i2. @⧣❨i1, 𝐭❨f❩❩ ≘ i2 → @⧣❨↑i1, 𝐭❨p⨮f❩❩ ≘ i2+p.
30 #p elim p -p /3 width=7 by pr_pat_push, pr_pat_next/
31 qed.
32
33 (*** at_plus2 *)
34 lemma tr_pat_plus_dx (f):
35       ∀i1,i,p,q. @⧣❨i1, 𝐭❨p⨮f❩❩ ≘ i → @⧣❨i1, 𝐭❨(p+q)⨮f❩❩ ≘ i+q.
36 #f #i1 #i #p #q #H elim q -q
37 /2 width=5 by pr_pat_next/
38 qed.
39
40 (* Inversions with pr_pat ***************************************************)
41
42 (*** at_inv_O1 *)
43 lemma tr_pat_inv_unit_sn (f):
44       ∀p,i2. @⧣❨𝟏, 𝐭❨p⨮f❩❩ ≘ i2 → p = i2.
45 #f #p elim p -p /2 width=6 by pr_pat_inv_unit_push/
46 #p #IH #i2 #H elim (pr_pat_inv_next … H) -H [|*: // ]
47 #j2 #Hj * -i2 /3 width=1 by eq_f/
48 qed-.
49
50 (*** at_inv_S1 *)
51 lemma tr_pat_inv_succ_sn (f):
52       ∀p,j1,i2. @⧣❨↑j1, 𝐭❨p⨮f❩❩ ≘ i2 →
53       ∃∃j2. @⧣❨j1, 𝐭❨f❩❩ ≘ j2 & j2+p = i2.
54 #f #p elim p -p /2 width=5 by pr_pat_inv_succ_push/
55 #p #IH #j1 #i2 #H elim (pr_pat_inv_next … H) -H [|*: // ]
56 #j2 #Hj * -i2 elim (IH … Hj) -IH -Hj
57 #i2 #Hi * -j2 /2 width=3 by ex2_intro/
58 qed-.
59
60 (* Destructions with pr_pat *************************************************)
61
62 (* Note: a better conclusion would be: "i1 + ↓p ≤ i2" *)
63 (*** at_increasing_plus *)
64 lemma tr_pat_increasing_plus (f):
65       ∀p,i1,i2. @⧣❨i1, 𝐭❨p⨮f❩❩ ≘ i2 → i1 + p ≤ ↑i2.
66 #f #p *
67 [ #i2 #H <(tr_pat_inv_unit_sn … H) -i2 //
68 | #i1 #i2 #H elim (tr_pat_inv_succ_sn … H) -H
69   #j1 #Ht * -i2 <pplus_succ_sn 
70   /4 width=2 by pr_pat_increasing, ple_plus_bi_dx, ple_succ_bi/
71 ]
72 qed-.
73
74 (*** at_fwd_id *)
75 lemma tr_pat_des_id (f):
76       ∀p,i. @⧣❨i, 𝐭❨p⨮f❩❩ ≘ i → 𝟏 = p.
77 #f #p #i #H lapply (pr_pat_des_id … H) -H #H
78 elim (eq_inv_pr_push_cons … H) -H //
79 qed-.