]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground/relocation/rtmap_istot.ma
propagating the arithmetics library, partial commit
[helm.git] / matita / matita / contribs / lambdadelta / ground / 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/notation/relations/ist_1.ma".
16 include "ground/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    'IsT 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 @(nat_ind_succ … n) -n //
48 #n #IH #f #Hf <tls_S
49 /3 width=1 by istot_tl/
50 qed.
51
52 (* Main forward lemmas on at ************************************************)
53
54 corec theorem at_ext: ∀f1,f2. 𝐓❪f1❫ → 𝐓❪f2❫ →
55                       (∀i,i1,i2. @❪i,f1❫ ≘ i1 → @❪i,f2❫ ≘ i2 → i1 = i2) →
56                       f1 ≡ f2.
57 #f1 cases (pn_split f1) * #g1 #H1
58 #f2 cases (pn_split f2) * #g2 #H2
59 #Hf1 #Hf2 #Hi
60 [ @(eq_push … H1 H2) @at_ext -at_ext /2 width=3 by istot_inv_push/ -Hf1 -Hf2
61   #i #i1 #i2 #Hg1 #Hg2 lapply (Hi (↑i) (↑i1) (↑i2) ??) /2 width=7 by at_push/
62 | cases (Hf2 (𝟏)) -Hf1 -Hf2 -at_ext
63   #j2 #Hf2 cases (at_increasing_strict … Hf2 … H2) -H2
64   lapply (Hi (𝟏) (𝟏) j2 … Hf2) /2 width=2 by at_refl/ -Hi -Hf2 -H1
65   #H2 #H cases (plt_ge_false … H) -H //
66 | cases (Hf1 (𝟏)) -Hf1 -Hf2 -at_ext
67   #j1 #Hf1 cases (at_increasing_strict … Hf1 … H1) -H1
68   lapply (Hi (𝟏) j1 (𝟏) Hf1 ?) /2 width=2 by at_refl/ -Hi -Hf1 -H2
69   #H1 #H cases (plt_ge_false … H) -H //
70 | @(eq_next … H1 H2) @at_ext -at_ext /2 width=3 by istot_inv_next/ -Hf1 -Hf2
71   #i #i1 #i2 #Hg1 #Hg2 lapply (Hi i (↑i1) (↑i2) ??) /2 width=5 by at_next/
72 ]
73 qed-.
74
75 (* Advanced properties on at ************************************************)
76
77 lemma at_dec: ∀f,i1,i2. 𝐓❪f❫ → Decidable (@❪i1,f❫ ≘ i2).
78 #f #i1 #i2 #Hf lapply (Hf i1) -Hf *
79 #j2 #Hf elim (eq_pnat_dec i2 j2)
80 [ #H destruct /2 width=1 by or_introl/
81 | /4 width=6 by at_mono, or_intror/
82 ]
83 qed-.
84
85 lemma is_at_dec: ∀f,i2. 𝐓❪f❫ → Decidable (∃i1. @❪i1,f❫ ≘ i2).
86 #f #i2 #Hf
87 lapply (dec_plt (λi1.@❪i1,f❫ ≘ i2) … (↑i2)) [| * ]
88 [ /2 width=1 by at_dec/
89 | * /3 width=2 by ex_intro, or_introl/
90 | #H @or_intror * #i1 #Hi12
91   /5 width=3 by at_increasing, plt_succ_dx, ex2_intro/
92 ]
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.