]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_istot.ma
rtmap (platform-indepent multple relocation): application and composition
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_istot.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_2/notation/relations/istotal_1.ma".
16 include "ground_2/relocation/rtmap_at.ma".
17
18 (* RELOCATION MAP ***********************************************************)
19
20 definition istot: predicate rtmap ≝ λf. ∀i. ∃j. @⦃i, f⦄ ≡ j.
21
22 interpretation "test for totality (rtmap)"
23    'IsTotal f = (istot f).
24
25 (* Basic inversion lemmas ***************************************************)
26
27 lemma istot_inv_push: ∀g. 𝐓⦃g⦄ → ∀f. ↑f = g → 𝐓⦃f⦄.
28 #g #Hg #f #H #i elim (Hg (⫯i)) -Hg
29 #j #Hg elim (at_inv_npx … Hg … H) -Hg -H /2 width=3 by ex_intro/
30 qed-.
31
32 lemma istot_inv_next: ∀g. 𝐓⦃g⦄ → ∀f. ⫯f = g → 𝐓⦃f⦄.
33 #g #Hg #f #H #i elim (Hg i) -Hg
34 #j #Hg elim (at_inv_xnx … Hg … H) -Hg -H /2 width=2 by ex_intro/
35 qed-.
36
37 (* Properties on tl *********************************************************)
38
39 lemma istot_tl: ∀f. 𝐓⦃f⦄ → 𝐓⦃⫱f⦄.
40 #f cases (pn_split f) *
41 #g * -f /2 width=3 by istot_inv_next, istot_inv_push/
42 qed.
43
44 (* Properties on tls ********************************************************)
45
46 lemma istot_tls: ∀n,f. 𝐓⦃f⦄ → 𝐓⦃⫱*[n]f⦄.
47 #n elim n -n /3 width=1 by istot_tl/
48 qed.
49
50 (* Advanced forward lemmas on at ********************************************)
51
52 let corec at_ext: ∀f1,f2. 𝐓⦃f1⦄ → 𝐓⦃f2⦄ →
53                   (∀i,i1,i2. @⦃i, f1⦄ ≡ i1 → @⦃i, f2⦄ ≡ i2 → i1 = i2) → f1 ≗ f2 ≝ ?.
54 #f1 cases (pn_split f1) * #g1 #H1
55 #f2 cases (pn_split f2) * #g2 #H2
56 #Hf1 #Hf2 #Hi
57 [ @(eq_push … H1 H2) @at_ext -at_ext /2 width=3 by istot_inv_push/ -Hf1 -Hf2
58   #i #i1 #i2 #Hg1 #Hg2 lapply (Hi (⫯i) (⫯i1) (⫯i2) ??) /2 width=7 by at_push/
59 | cases (Hf2 0) -Hf1 -Hf2 -at_ext
60   #j2 #Hf2 cases (at_increasing_strict … Hf2 … H2) -H2
61   lapply (Hi 0 0 j2 … Hf2) /2 width=2 by at_refl/ -Hi -Hf2 -H1
62   #H2 #H cases (lt_le_false … H) -H //
63 | cases (Hf1 0) -Hf1 -Hf2 -at_ext
64   #j1 #Hf1 cases (at_increasing_strict … Hf1 … H1) -H1
65   lapply (Hi 0 j1 0 Hf1 ?) /2 width=2 by at_refl/ -Hi -Hf1 -H2
66   #H1 #H cases (lt_le_false … H) -H //
67 | @(eq_next … H1 H2) @at_ext -at_ext /2 width=3 by istot_inv_next/ -Hf1 -Hf2
68   #i #i1 #i2 #Hg1 #Hg2 lapply (Hi i (⫯i1) (⫯i2) ??) /2 width=5 by at_next/
69 ]
70 qed-.
71
72 (* Main properties on at ****************************************************)
73
74 lemma at_dec: ∀f,i1,i2. 𝐓⦃f⦄ → Decidable (@⦃i1, f⦄ ≡ i2).
75 #f #i1 #i2 #Hf lapply (Hf i1) -Hf *
76 #j2 #Hf elim (eq_nat_dec i2 j2)
77 [ #H destruct /2 width=1 by or_introl/
78 | /4 width=6 by at_mono, or_intror/
79 ]
80 qed-.
81
82 lemma is_at_dec_le: ∀f,i2,i. 𝐓⦃f⦄ → (∀i1. i1 + i ≤ i2 → @⦃i1, f⦄ ≡ i2 → ⊥) → Decidable (∃i1. @⦃i1, f⦄ ≡ i2).
83 #f #i2 #i #Hf elim i -i
84 [ #Ht @or_intror * /3 width=3 by at_increasing/
85 | #i #IH #Ht elim (at_dec f (i2-i) i2) /3 width=2 by ex_intro, or_introl/
86   #Hi2 @IH -IH #i1 #H #Hi elim (le_to_or_lt_eq … H) -H /2 width=3 by/
87   #H destruct -Ht /2 width=1 by/
88 ]
89 qed-.
90
91 lemma is_at_dec: ∀f,i2. 𝐓⦃f⦄ → Decidable (∃i1. @⦃i1, f⦄ ≡ i2).
92 #f #i2 #Hf @(is_at_dec_le ?? (⫯i2)) /2 width=4 by lt_le_false/
93 qed-.
94
95 (* Advanced properties on isid **********************************************)
96
97 lemma isid_at_total: ∀f. 𝐓⦃f⦄ → (∀i1,i2. @⦃i1, f⦄ ≡ i2 → i1 = i2) → 𝐈⦃f⦄.
98 #f #H1f #H2f @isid_at
99 #i lapply (H1f i) -H1f *
100 #j #Hf >(H2f … Hf) in ⊢ (???%); -H2f //
101 qed.