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 "basic_2/relocation/lleq_lleq.ma".
16 include "basic_2/computation/lpxs.ma".
17 include "basic_2/computation/fpns.ma".
19 (* COMPUTATION FOR "BIG TREE" NORMAL FORMS **********************************)
21 (* Advanced inversion lemmas ************************************************)
23 lemma fpns_inv_gen: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊢ ⋕➡*[h, g] ⦃G2, L2, T2⦄ →
24 ∧∧ G1 = G2 & ⦃G1, L1⦄ ⊢ ➡*[h, g] L2 & L1 ⋕[T1] L2 & T1 = T2.
25 #h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpns_ind … H) -G2 -L2 -T2 /2 width=1 by and4_intro/
26 #G #G2 #L #L2 #T #T2 #_ #H2 * #HG #HL1 #HT1 #HT destruct
27 elim (fpn_inv_gen … H2) -H2 #HG #HL2 #HT #HT2 destruct
28 /3 width=3 by lpxs_strap1, lleq_trans, and4_intro/
31 (* Advanced properties ******************************************************)
33 lemma lpxs_lleq_fpns: ∀h,g,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → L1 ⋕[T] L2 →
34 ⦃G, L1, T⦄ ⊢ ⋕➡*[h, g] ⦃G, L2, T⦄.
35 #h #g #G #L1 #L2 #T #H @(lpxs_ind … H) -L2 //
36 #L #L2 #HL1 #HL2 #IHL1 #HL12 elim (lleq_dec T L1 L) #HT
37 [ -HL1 @fpns_strap1 [4: @IHL1 // |1,2,3: skip ]
38 /3 width=3 by fpn_intro, lleq_canc_sn/
41 (* Main properties **********************************************************)
43 theorem fpns_trans: tri_transitive … fqus.
44 /2 width=5 by tri_TC_transitive/ qed-.
48 lemma fpns_refl: ∀h,g. tri_reflexive … (fpns h g).
49 /2 width=1 by and3_intro/ qed.
51 lemma fpn_fpns: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊢ ➡[h, g] ⦃G2, L2, T2⦄ →
52 ⦃G1, L1, T1⦄ ⊢ ➡*[h, g] ⦃G2, L2, T2⦄.
53 #h #g #G1 #G2 #L1 #L2 #T1 #T2 * /3 width=1 by lpx_lpxs, and3_intro/
56 lemma fpns_strap1: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊢ ➡*[h, g] ⦃G, L, T⦄ →
57 ⦃G, L, T⦄ ⊢ ➡[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊢ ➡*[h, g] ⦃G2, L2, T2⦄.
58 #h #g #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 * #H1G #H1L #G1T *
59 /3 width=3 by lpxs_strap1, and3_intro/
62 lemma fpns_strap2: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊢ ➡[h, g] ⦃G, L, T⦄ →
63 ⦃G, L, T⦄ ⊢ ➡*[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊢ ➡*[h, g] ⦃G2, L2, T2⦄.
64 #h #g #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 * #H1G #H1L #G1T *
65 /3 width=3 by lpxs_strap2, and3_intro/
68 (* Basic forward lemmas *****************************************************)
70 lemma fpns_fwd_bteq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊢ ➡*[h, g] ⦃G2, L2, T2⦄ →
71 ⦃G1, L1, T1⦄ ⋕ ⦃G2, L2, T2⦄.
72 #h #g #G1 #G2 #L1 #L2 #T1 #T2 * /3 width=4 by lpxs_fwd_length, and3_intro/