1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "ground/arith/pnat_pred.ma".
16 include "ground/arith/pnat_lt.ma".
17 include "ground/relocation/pr_pat.ma".
19 (* POSITIVE APPLICATION FOR PARTIAL RELOCATION MAPS *************************)
21 (* Destructions with plt and ple ********************************************)
24 lemma pr_pat_increasing (i2) (i1) (f):
25 @❨i1,f❩ ≘ i2 → i1 ≤ i2.
27 [ #i1 #f #Hf elim (pr_pat_inv_unit_dx … Hf) -Hf //
29 #i1 #f #Hf elim (pr_pat_inv_succ_bi … Hf) -Hf [1,4: * |*: // ]
30 /3 width=2 by ple_succ_bi, ple_succ_dx/
34 (*** at_increasing_strict *)
35 lemma pr_pat_increasing_strict (g) (i1) (i2):
36 @❨i1,g❩ ≘ i2 → ∀f. ↑f = g →
37 ∧∧ i1 < i2 & @❨i1,f❩ ≘ ↓i2.
38 #g #i1 #i2 #Hg #f #H elim (pr_pat_inv_next … Hg … H) -Hg -H
39 /4 width=2 by conj, pr_pat_increasing, ple_succ_bi/
43 lemma pr_pat_des_id (f) (i): @❨i,f❩ ≘ i → ⫯⫰f = f.
44 #f elim (pr_map_split_tl f) //
45 #H #i #Hf elim (pr_pat_inv_next … Hf … H) -Hf -H
46 #j2 #Hg #H destruct lapply (pr_pat_increasing … Hg) -Hg
47 #H elim (plt_ge_false … H) -H //
50 (* Constructions with ple ***************************************************)
53 lemma pr_pat_le_ex (j2) (i2) (f):
54 @❨i2,f❩ ≘ j2 → ∀i1. i1 ≤ i2 →
55 ∃∃j1. @❨i1,f❩ ≘ j1 & j1 ≤ j2.
56 #j2 elim j2 -j2 [2: #j2 #IH ] #i2 #f #Hf
57 [ elim (pr_pat_inv_succ_dx … Hf) -Hf [1,3: * |*: // ]
58 #g [ #x2 ] #Hg [ #H2 ] #H0
59 [ * /3 width=3 by pr_pat_refl, ex2_intro/
60 #i1 #Hi12 destruct lapply (ple_inv_succ_bi … Hi12) -Hi12
61 #Hi12 elim (IH … Hg … Hi12) -x2 -IH
62 /3 width=7 by pr_pat_push, ex2_intro, ple_succ_bi/
63 | #i1 #Hi12 elim (IH … Hg … Hi12) -IH -i2
64 /3 width=5 by pr_pat_next, ex2_intro, ple_succ_bi/
66 | elim (pr_pat_inv_unit_dx … Hf) -Hf //
67 #g * -i2 #H2 #i1 #Hi12 <(ple_inv_unit_dx … Hi12)
68 /3 width=3 by pr_pat_refl, ex2_intro/
73 lemma pr_pat_id_le (i1) (i2):
74 i1 ≤ i2 → ∀f. @❨i2,f❩ ≘ i2 → @❨i1,f❩ ≘ i1.
76 @(ple_ind_alt … H) -i1 -i2 [ #i2 | #i1 #i2 #_ #IH ] #f #Hf
77 lapply (pr_pat_des_id … Hf) #H <H in Hf; -H
78 /4 width=7 by pr_pat_inv_succ_push_succ, pr_pat_push, pr_pat_refl/