]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_cpxs.ma
milestone in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / csx_cpxs.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 "basic_2/rt_computation/cpxs_tdeq.ma".
16 include "basic_2/rt_computation/cpxs_cpxs.ma".
17 include "basic_2/rt_computation/csx_csx.ma".
18
19 (* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************)
20
21 (* Properties with unbound context-sensitive rt-computation for terms *******)
22
23 (* Basic_1: was just: sn3_intro *)
24 lemma csx_intro_cpxs: ∀h,G,L,T1.
25                       (∀T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → (T1 ≛ T2 → ⊥) → ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃T2⦄) →
26                       ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃T1⦄.
27 /4 width=1 by cpx_cpxs, csx_intro/ qed-.
28
29 (* Basic_1: was just: sn3_pr3_trans *)
30 lemma csx_cpxs_trans: ∀h,G,L,T1. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃T1⦄ →
31                       ∀T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃T2⦄.
32 #h #G #L #T1 #HT1 #T2 #H @(cpxs_ind … H) -T2
33 /2 width=3 by csx_cpx_trans/
34 qed-.
35
36 (* Eliminators with unbound context-sensitive rt-computation for terms ******)
37
38 lemma csx_ind_cpxs_tdeq: ∀h,G,L. ∀Q:predicate term.
39                          (∀T1. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃T1⦄ →
40                                (∀T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → (T1 ≛ T2 → ⊥) → Q T2) → Q T1
41                          ) →
42                          ∀T1. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃T1⦄ →
43                          ∀T0. ⦃G, L⦄ ⊢ T1 ⬈*[h] T0 → ∀T2. T0 ≛ T2 → Q T2.
44 #h #G #L #Q #IH #T1 #H @(csx_ind … H) -T1
45 #T1 #HT1 #IH1 #T0 #HT10 #T2 #HT02
46 @IH -IH /3 width=3 by csx_cpxs_trans, csx_tdeq_trans/ -HT1 #V2 #HTV2 #HnTV2
47 lapply (tdeq_tdneq_trans … HT02 … HnTV2) -HnTV2 #H
48 elim (tdeq_cpxs_trans … HT02 … HTV2) -T2 #V0 #HTV0 #HV02
49 lapply (tdneq_tdeq_canc_dx … H … HV02) -H #HnTV0
50 elim (tdeq_dec T1 T0) #H
51 [ lapply (tdeq_tdneq_trans … H … HnTV0) -H -HnTV0 #Hn10
52   lapply (cpxs_trans … HT10 … HTV0) -T0 #H10
53   elim (cpxs_tdneq_fwd_step_sn … H10 …  Hn10) -H10 -Hn10
54   /3 width=8 by tdeq_trans/
55 | elim (cpxs_tdneq_fwd_step_sn … HT10 … H) -HT10 -H #T #V #HT1 #HnT1 #HTV #HVT0
56   elim (tdeq_cpxs_trans … HVT0 … HTV0) -T0
57   /3 width=8 by cpxs_trans, tdeq_trans/
58 ]
59 qed-.
60
61 (* Basic_2A1: was: csx_ind_alt *)
62 lemma csx_ind_cpxs: ∀h,G,L. ∀Q:predicate term.
63                     (∀T1. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃T1⦄ →
64                           (∀T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → (T1 ≛ T2 → ⊥) → Q T2) → Q T1
65                     ) →
66                     ∀T. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃T⦄ →  Q T.
67 #h #G #L #Q #IH #T #HT
68 @(csx_ind_cpxs_tdeq … IH … HT) -IH -HT // (**) (* full auto fails *)
69 qed-.