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 notation "hvbox( ⦃ term 46 L1, break term 46 T1 ⦄ ⧁ break ⦃ term 46 L2 , break term 46 T2 ⦄ )"
16 non associative with precedence 45
17 for @{ 'RestSupTerm $L1 $T1 $L2 $T2 }.
19 include "basic_2/grammar/cl_weight.ma".
20 include "basic_2/substitution/lift.ma".
22 (* RESTRICTED SUPCLOSURE ****************************************************)
24 inductive frsup: bi_relation lenv term ≝
25 | frsup_bind_sn: ∀a,I,L,V,T. frsup L (ⓑ{a,I}V.T) L V
26 | frsup_bind_dx: ∀a,I,L,V,T. frsup L (ⓑ{a,I}V.T) (L.ⓑ{I}V) T
27 | frsup_flat_sn: ∀I,L,V,T. frsup L (ⓕ{I}V.T) L V
28 | frsup_flat_dx: ∀I,L,V,T. frsup L (ⓕ{I}V.T) L T
32 "restricted structural predecessor (closure)"
33 'RestSupTerm L1 T1 L2 T2 = (frsup L1 T1 L2 T2).
35 (* Basic inversion lemmas ***************************************************)
37 fact frsup_inv_atom1_aux: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⧁ ⦃L2, T2⦄ →
39 #L1 #L2 #T1 #T2 * -L1 -L2 -T1 -T2
40 [ #a #I #L #V #T #J #H destruct
41 | #a #I #L #V #T #J #H destruct
42 | #I #L #V #T #J #H destruct
43 | #I #L #V #T #J #H destruct
47 lemma frsup_inv_atom1: ∀J,L1,L2,T2. ⦃L1, ⓪{J}⦄ ⧁ ⦃L2, T2⦄ → ⊥.
48 /2 width=7 by frsup_inv_atom1_aux/ qed-.
50 fact frsup_inv_bind1_aux: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⧁ ⦃L2, T2⦄ →
51 ∀b,J,W,U. T1 = ⓑ{b,J}W.U →
53 (L2 = L1.ⓑ{J}W ∧ T2 = U).
54 #L1 #L2 #T1 #T2 * -L1 -L2 -T1 -T2
55 [ #a #I #L #V #T #b #J #W #U #H destruct /3 width=1/
56 | #a #I #L #V #T #b #J #W #U #H destruct /3 width=1/
57 | #I #L #V #T #b #J #W #U #H destruct
58 | #I #L #V #T #b #J #W #U #H destruct
62 lemma frsup_inv_bind1: ∀b,J,L1,L2,W,U,T2. ⦃L1, ⓑ{b,J}W.U⦄ ⧁ ⦃L2, T2⦄ →
64 (L2 = L1.ⓑ{J}W ∧ T2 = U).
65 /2 width=4 by frsup_inv_bind1_aux/ qed-.
67 fact frsup_inv_flat1_aux: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⧁ ⦃L2, T2⦄ →
68 ∀J,W,U. T1 = ⓕ{J}W.U →
69 L2 = L1 ∧ (T2 = W ∨ T2 = U).
70 #L1 #L2 #T1 #T2 * -L1 -L2 -T1 -T2
71 [ #a #I #L #V #T #J #W #U #H destruct
72 | #a #I #L #V #T #J #W #U #H destruct
73 | #I #L #V #T #J #W #U #H destruct /3 width=1/
74 | #I #L #V #T #J #W #U #H destruct /3 width=1/
78 lemma frsup_inv_flat1: ∀J,L1,L2,W,U,T2. ⦃L1, ⓕ{J}W.U⦄ ⧁ ⦃L2, T2⦄ →
79 L2 = L1 ∧ (T2 = W ∨ T2 = U).
80 /2 width=4 by frsup_inv_flat1_aux/ qed-.
82 (* Basic forward lemmas *****************************************************)
84 lemma frsup_fwd_fw: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⧁ ⦃L2, T2⦄ → ♯{L2, T2} < ♯{L1, T1}.
85 #L1 #L2 #T1 #T2 * -L1 -L2 -T1 -T2 /width=1/
88 lemma frsup_fwd_lw: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⧁ ⦃L2, T2⦄ → ♯{L1} ≤ ♯{L2}.
89 #L1 #L2 #T1 #T2 * -L1 -L2 -T1 -T2 /width=1/
92 lemma frsup_fwd_tw: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⧁ ⦃L2, T2⦄ → ♯{T2} < ♯{T1}.
93 #L1 #L2 #T1 #T2 * -L1 -L2 -T1 -T2 /width=1/ /2 width=1 by le_minus_to_plus/
96 lemma frsup_fwd_append: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⧁ ⦃L2, T2⦄ → ∃L. L2 = L1 @@ L.
97 #L1 #L2 #T1 #T2 * -L1 -L2 -T1 -T2
99 | #a #I #L #V #_ @(ex_intro … (⋆.ⓑ{I}V)) //
101 #I #L #V #T @(ex_intro … (⋆)) //
104 (* Advanced forward lemmas **************************************************)
106 lemma lift_frsup_trans: ∀T1,U1,d,e. ⇧[d, e] T1 ≡ U1 →
107 ∀L,K,U2. ⦃L, U1⦄ ⧁ ⦃L @@ K, U2⦄ →
108 ∃T2. ⇧[d + |K|, e] T2 ≡ U2.
109 #T1 #U1 #d #e * -T1 -U1 -d -e
110 [5: #a #I #V1 #W1 #T1 #U1 #d #e #HVW1 #HTU1 #L #K #X #H
111 elim (frsup_inv_bind1 … H) -H *
112 [ -HTU1 #H1 #H2 destruct
113 >(append_inv_refl_dx … H1) -L -K normalize /2 width=2/
114 | -HVW1 #H1 #H2 destruct
115 >(append_inv_pair_dx … H1) -L -K normalize /2 width=2/
117 |6: #I #V1 #W1 #T1 #U1 #d #e #HVW1 #HUT1 #L #K #X #H
118 elim (frsup_inv_flat1 … H) -H #H1 * #H2 destruct
119 >(append_inv_refl_dx … H1) -L -K normalize /2 width=2/
121 #i #d #e [2,3: #_ ] #L #K #X #H
122 elim (frsup_inv_atom1 … H)