definition dsubstable: predicate (relation term) ≝ λR.
∀D1,D2. R D1 D2 → ∀M1,M2. R M1 M2 → ∀d. R ([d ⬐ D1] M1) ([d ⬐ D2] M2).
+
+lemma star_dsubstable_dx: ∀R. dsubstable_dx R → dsubstable_dx (star … R).
+#R #HR #D #M1 #M2 #H elim H -M2 // /3 width=3/
+qed.
+
+lemma star_dsubstable_sn: ∀R. dsubstable_sn R → dsubstable_sn (star … R).
+#R #HR #D1 #D2 #H elim H -D2 // /3 width=3/
+qed.
+
+lemma lstar_dsubstable_dx: ∀T,R. (∀t. dsubstable_dx (R t)) →
+ ∀l. dsubstable_dx (lstar T … R l).
+#T #R #HR #l #D #M1 #M2 #H elim H -l -M1 -M2 // /3 width=3/
+qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "hap_reduction.ma".
-
-(* KASHIMA'S "HAP" COMPUTATION **********************************************)
-
-(* Note: this is the "head in application" computation of:
- R. Kashima: "A proof of the Standization Theorem in λ-Calculus". Typescript note, (2000).
-*)
-definition hap: relation term ≝ star … hap1.
-
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "labelled_sequential_reduction.ma".
-
-(* KASHIMA'S "HAP" COMPUTATION (SINGLE STEP) ********************************)
-
-(* Note: this is one step of the "head in application" computation of:
- R. Kashima: "A proof of the Standization Theorem in λ-Calculus". Typescript note, (2000).
-*)
-inductive hap1: relation term ≝
-| hap1_beta: ∀B,A. hap1 (@B.𝛌.A) ([⬐B]A)
-| hap1_appl: ∀B,A1,A2. hap1 A1 A2 → hap1 (@B.A1) (@B.A2)
-.
-
-lemma hap1_lift: liftable hap1.
-#h #M1 #M2 #H elim H -M1 -M2 normalize /2 width=1/
-#B #A #d <dsubst_lift_le //
-qed.
-
-lemma hap1_inv_lift: deliftable hap1.
-#h #N1 #N2 #H elim H -N1 -N2
-[ #D #C #d #M1 #H
- elim (lift_inv_appl … H) -H #B #M #H0 #HM #H destruct
- elim (lift_inv_abst … HM) -HM #A #H0 #H destruct /3 width=3/
-| #D1 #C1 #C2 #_ #IHC12 #d #M1 #H
- elim (lift_inv_appl … H) -H #B #A1 #H1 #H2 #H destruct
- elim (IHC12 ???) -IHC12 [4: // |2,3: skip ] #A2 #HA12 #H destruct (**) (* simplify line *)
- @(ex2_1_intro … (@B.A2)) // /2 width=1/
-]
-qed-.
-
-lemma hap1_dsubstable: dsubstable_dx hap1.
-#D1 #M1 #M2 #H elim H -M1 -M2 normalize /2 width=1/
-#D2 #A #d >dsubst_dsubst_ge //
-qed.
-
-lemma hap1_lsred: ∀M,N. hap1 M N →
- ∃∃p. in_spine p & M ⇀[p] N.
-#M #N #H elim H -M -N /2 width=3/
-#B #A1 #A2 #_ * /3 width=3/
-qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "labelled_sequential_computation.ma".
+include "labelled_hap_reduction.ma".
+
+(* KASHIMA'S "HAP" COMPUTATION (LABELLED MULTISTEP) *************************)
+
+(* Note: this is the "head in application" computation of:
+ R. Kashima: "A proof of the Standization Theorem in λ-Calculus". Typescript note, (2000).
+*)
+definition lhap: rpseq → relation term ≝ lstar … lhap1.
+
+interpretation "labelled 'hap' computation"
+ 'HApStar M p N = (lhap p M N).
+
+notation "hvbox( M break ⓗ⇀* [ term 46 p ] break term 46 N )"
+ non associative with precedence 45
+ for @{ 'HApStar $M $p $N }.
+
+lemma lhap1_lhap: ∀p,M1,M2. M1 ⓗ⇀[p] M2 → M1 ⓗ⇀*[p::◊] M2.
+/2 width=1/
+qed.
+
+lemma lhap_inv_nil: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → ◊ = s → M1 = M2.
+/2 width=5 by lstar_inv_nil/
+qed-.
+
+lemma lhap_inv_cons: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → ∀q,r. q::r = s →
+ ∃∃M. M1 ⓗ⇀[q] M & M ⓗ⇀*[r] M2.
+/2 width=3 by lstar_inv_cons/
+qed-.
+
+lemma lhap_inv_lhap1: ∀p,M1,M2. M1 ⓗ⇀*[p::◊] M2 → M1 ⓗ⇀[p] M2.
+/2 width=1 by lstar_inv_step/
+qed-.
+
+lemma lhap_lift: ∀s. liftable (lhap s).
+/2 width=1/
+qed.
+
+lemma lhap_inv_lift: ∀s. deliftable_sn (lhap s).
+/3 width=3 by lstar_deliftable_sn, lhap1_inv_lift/
+qed-.
+
+lemma lhap_dsubst: ∀s. dsubstable_dx (lhap s).
+/2 width=1/
+qed.
+(*
+theorem lhap_mono: ∀s. singlevalued … (lhap s).
+/3 width=7 by lstar_singlevalued, lhap1_mono/
+qed-.
+*)
+theorem lhap_trans: ∀s1,M1,M. M1 ⓗ⇀*[s1] M →
+ ∀s2,M2. M ⓗ⇀*[s2] M2 → M1 ⓗ⇀*[s1@s2] M2.
+/2 width=3 by lstar_trans/
+qed-.
+(*
+lemma hap_appl: appl_compatible_dx hap.
+/3 width=1/
+qed.
+*)
+lemma lhap_spine_fwd: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → is_spine s.
+#s #M1 #M2 #H elim H -s -M1 -M2 //
+#p #M1 #M #HM1 #s #M2 #_ #IHM2
+lapply (lhap1_spine_fwd … HM1) -HM1 /2 width=1/
+qed-.
+
+lemma lhap_lsreds_fwd: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → M1 ⇀*[s] M2.
+#s #M1 #M2 #H elim H -s -M1 -M2 //
+#p #M1 #M #HM1 #s #M2 #_ #IHM2
+lapply (lhap1_lsred_fwd … HM1) -HM1 /2 width=3/
+qed-.
+
+lemma lhap_le_fwd: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → is_le s.
+(*
+#M1 #M2 #H @(star_ind_l ??????? H) -M1 /3 width=3/
+#M1 #M #HM1 #H * #s * #H1s #H2s
+generalize in match HM1; -HM1 generalize in match M1; -M1
+@(star_ind_l ??????? H) -M
+[ #M1 #HM12 elim (hap1_lsred … HM12) -HM12 /4 width=3/
+| #M0 #M1 #HM01 #HM12 #_ #M #HM0 #HM02
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "labelled_sequential_reduction.ma".
+
+(* KASHIMA'S "HAP" COMPUTATION (LABELLED SINGLE STEP) ***********************)
+
+(* Note: this is one step of the "head in application" computation of:
+ R. Kashima: "A proof of the Standization Theorem in λ-Calculus". Typescript note, (2000).
+*)
+inductive lhap1: rptr → relation term ≝
+| hap1_beta: ∀B,A. lhap1 (◊) (@B.𝛌.A) ([⬐B]A)
+| hap1_appl: ∀p,B,A1,A2. lhap1 p A1 A2 → lhap1 (false::p) (@B.A1) (@B.A2)
+.
+
+interpretation "labelled 'hap' reduction"
+ 'HAp M p N = (lhap1 p M N).
+
+notation "hvbox( M break ⓗ⇀ [ term 46 p ] break term 46 N )"
+ non associative with precedence 45
+ for @{ 'HAp $M $p $N }.
+
+lemma lhap1_inv_abst_sn: ∀p,M,N. M ⓗ⇀[p] N → ∀A. 𝛌.A = M → ⊥.
+#p #M #N * -p -M -N
+[ #B #A #A0 #H destruct
+| #p #B #A1 #A2 #_ #A0 #H destruct
+]
+qed-.
+
+lemma lhap1_inv_appl_sn: ∀p,M,N. M ⓗ⇀[p] N → ∀B,A. @B.A = M →
+ (∃∃C. ◊ = p & 𝛌.C = A & [⬐B]C = N) ∨
+ ∃∃q,C. A ⓗ⇀[q] C & false::q = p & @B.C = N.
+#p #M #N * -p -M -N
+[ #B #A #B0 #A0 #H destruct /3 width=3/
+| #p #B #A1 #A2 #HA12 #B0 #A0 #H destruct /3 width=5/
+]
+qed-.
+
+lemma lhap1_inv_abst_dx: ∀p,M,N. M ⓗ⇀[p] N → ∀C. 𝛌.C = N →
+ ∃∃B,A. ◊ = p & @B.𝛌.A = M & 𝛌.C = [⬐B]A.
+#p #M #N * -p -M -N
+[ #B #A #C #H /2 width=4/
+| #p #B #A1 #A2 #_ #C #H destruct
+]
+qed-.
+
+lemma lhap1_lift: ∀p. liftable (lhap1 p).
+#p #h #M1 #M2 #H elim H -p -M1 -M2 normalize /2 width=1/
+#B #A #d <dsubst_lift_le //
+qed.
+
+lemma lhap1_inv_lift: ∀p. deliftable_sn (lhap1 p).
+#p #h #N1 #N2 #H elim H -p -N1 -N2
+[ #D #C #d #M1 #H
+ elim (lift_inv_appl … H) -H #B #M #H0 #HM #H destruct
+ elim (lift_inv_abst … HM) -HM #A #H0 #H destruct /3 width=3/
+| #p #D1 #C1 #C2 #_ #IHC12 #d #M1 #H
+ elim (lift_inv_appl … H) -H #B #A1 #H1 #H2 #H destruct
+ elim (IHC12 ???) -IHC12 [4: // |2,3: skip ] #A2 #HA12 #H destruct (**) (* simplify line *)
+ @(ex2_intro … (@B.A2)) // /2 width=1/
+]
+qed-.
+
+lemma lhap1_dsubst: ∀p. dsubstable_dx (lhap1 p).
+#p #D1 #M1 #M2 #H elim H -p -M1 -M2 normalize /2 width=1/
+#D2 #A #d >dsubst_dsubst_ge //
+qed.
+
+lemma lhap1_spine_fwd: ∀p,M,N. M ⓗ⇀[p] N → in_spine p.
+#p #M #N #H elim H -p -M -N // /2 width=1/
+qed-.
+
+lemma lhap1_lsred_fwd: ∀p,M,N. M ⓗ⇀[p] N → M ⇀[p] N.
+#p #M #N #H elim H -p -M -N // /2 width=1/
+qed-.
+
+lemma lhap1_le_fwd: ∀p1,M1,M. M1 ⓗ⇀[p1] M → ∀p2,M2. M ⓗ⇀[p2] M2 → p1 ≤ p2.
+#p1 #M1 #M #H elim H -p1 -M1 -M //
+#p1 #B #A1 #A2 #HA12 #IHA12 #p2 #M2 #H
+elim (lhap1_inv_appl_sn … H ???) -H [5: // |2,3: skip ] * (**) (* simplify line *)
+[ -IHA12 #C2 #Hp2 #HAC2 #_
+ elim (lhap1_inv_abst_dx … HA12 … HAC2) -A2 #B1 #C1 #Hp1 #_ #_ //
+| -HA12 /3 width=2/
+]
+qed-.
(* LABELLED SEQUENTIAL COMPUTATION (MULTISTEP) ******************************)
-(* Note: this comes from "star term lsred" *)
-inductive lsreds: rpseq → relation term ≝
-| lsreds_nil : ∀M. lsreds (◊) M M
-| lsreds_cons: ∀p,M1,M. M1 ⇀[p] M →
- ∀s,M2. lsreds s M M2 → lsreds (p::s) M1 M2
-.
+definition lsreds: rpseq → relation term ≝ lstar … lsred.
interpretation "labelled sequential computation"
'SeqRedStar M s N = (lsreds s M N).
for @{ 'SeqRedStar $M $s $N }.
lemma lsred_lsreds: ∀p,M1,M2. M1 ⇀[p] M2 → M1 ⇀*[p::◊] M2.
-/2 width=3/
+/2 width=1/
qed.
lemma lsreds_inv_nil: ∀s,M1,M2. M1 ⇀*[s] M2 → ◊ = s → M1 = M2.
-#s #M1 #M2 * -s -M1 -M2 //
-#p #M1 #M #_ #s #M2 #_ #H destruct
+/2 width=5 by lstar_inv_nil/
qed-.
lemma lsreds_inv_cons: ∀s,M1,M2. M1 ⇀*[s] M2 → ∀q,r. q::r = s →
∃∃M. M1 ⇀[q] M & M ⇀*[r] M2.
-#s #M1 #M2 * -s -M1 -M2
-[ #M #q #r #H destruct
-| #p #M1 #M #HM1 #s #M2 #HM2 #q #r #H destruct /2 width=3/
-]
+/2 width=3 by lstar_inv_cons/
qed-.
lemma lsreds_inv_lsred: ∀p,M1,M2. M1 ⇀*[p::◊] M2 → M1 ⇀[p] M2.
-#p #M1 #M2 #H
-elim (lsreds_inv_cons … H ???) -H [4: // |2,3: skip ] #M #HM1 #H (**) (* simplify line *)
-<(lsreds_inv_nil … H ?) -H //
-qed-.
-
-(* Note: "|s|" should be unparetesized *)
-lemma lsreds_fwd_mult: ∀s,M1,M2. M1 ⇀*[s] M2 → #{M2} ≤ #{M1} ^ (2 ^ (|s|)).
-#s #M1 #M2 #H elim H -s -M1 -M2 normalize //
-#p #M1 #M #HM1 #s #M2 #_ #IHM2
-lapply (lsred_fwd_mult … HM1) -p #HM1
-@(transitive_le … IHM2) -M2
-/3 width=1 by le_exp1, lt_O_exp, lt_to_le/ (**) (* auto: slow without trace *)
+/2 width=1 by lstar_inv_step/
qed-.
lemma lsreds_lift: ∀s. liftable (lsreds s).
-#s #h #M1 #M2 #H elim H -s -M1 -M2 // /3 width=3/
+/2 width=1/
qed.
-lemma lsreds_inv_lift: ∀s. deliftable (lsreds s).
-#s #h #N1 #N2 #H elim H -s -N1 -N2 /2 width=3/
-#p #N1 #N #HN1 #s #N2 #_ #IHN2 #d #M1 #HMN1
-elim (lsred_inv_lift … HN1 … HMN1) -N1 #M #HM1 #HMN
-elim (IHN2 … HMN) -N /3 width=3/
+lemma lsreds_inv_lift: ∀s. deliftable_sn (lsreds s).
+/3 width=3 by lstar_deliftable_sn, lsred_inv_lift/
qed-.
lemma lsreds_dsubst: ∀s. dsubstable_dx (lsreds s).
-#s #D #M1 #M2 #H elim H -s -M1 -M2 // /3 width=3/
+/2 width=1/
qed.
theorem lsreds_mono: ∀s. singlevalued … (lsreds s).
-#s #M #N1 #H elim H -s -M -N1
-[ /2 width=3 by lsreds_inv_nil/
-| #p #M #M1 #HM1 #s #N1 #_ #IHMN1 #N2 #H
- elim (lsreds_inv_cons … H ???) -H [4: // |2,3: skip ] #M2 #HM2 #HMN2 (**) (* simplify line *)
- lapply (lsred_mono … HM1 … HM2) -M #H destruct /2 width=1/
-]
+/3 width=7 by lstar_singlevalued, lsred_mono/
qed-.
theorem lsreds_trans: ∀s1,M1,M. M1 ⇀*[s1] M →
∀s2,M2. M ⇀*[s2] M2 → M1 ⇀*[s1@s2] M2.
-#s1 #M1 #M #H elim H -s1 -M1 -M normalize // /3 width=3/
+/2 width=3 by lstar_trans/
+qed-.
+
+(* Note: "|s|" should be unparetesized *)
+lemma lsreds_fwd_mult: ∀s,M1,M2. M1 ⇀*[s] M2 → #{M2} ≤ #{M1} ^ (2 ^ (|s|)).
+#s #M1 #M2 #H elim H -s -M1 -M2 normalize //
+#p #M1 #M #HM1 #s #M2 #_ #IHM2
+lapply (lsred_fwd_mult … HM1) -p #HM1
+@(transitive_le … IHM2) -M2
+/3 width=1 by le_exp1, lt_O_exp, lt_to_le/ (**) (* auto: slow without trace *)
qed-.
#B #A #d <dsubst_lift_le //
qed.
-lemma lsred_inv_lift: ∀p. deliftable (lsred p).
+lemma lsred_inv_lift: ∀p. deliftable_sn (lsred p).
#p #h #N1 #N2 #H elim H -p -N1 -N2
[ #D #C #d #M1 #H
elim (lift_inv_appl … H) -H #B #M #H0 #HM #H destruct
elim (lift_inv_abst … HM) -HM #A #H0 #H destruct /3 width=3/
| #p #C1 #C2 #_ #IHC12 #d #M1 #H
- elim (lift_inv_abst … H) -H #A1 #H0 #H destruct
- elim (IHC12 ???) -IHC12 [4: // |2,3: skip ] #A2 #HA12 #H destruct (**) (* simplify line *)
- @(ex2_1_intro … (𝛌.A2)) // /2 width=1/
+ elim (lift_inv_abst … H) -H #A1 #HAC1 #H
+ elim (IHC12 … HAC1) -C1 #A2 #HA12 #HAC2 destruct
+ @(ex2_intro … (𝛌.A2)) // /2 width=1/
| #p #D1 #D2 #C1 #_ #IHD12 #d #M1 #H
- elim (lift_inv_appl … H) -H #B1 #A #H1 #H2 #H destruct
- elim (IHD12 ???) -IHD12 [4: // |2,3: skip ] #B2 #HB12 #H destruct (**) (* simplify line *)
- @(ex2_1_intro … (@B2.A)) // /2 width=1/
+ elim (lift_inv_appl … H) -H #B1 #A #HBD1 #H1 #H2
+ elim (IHD12 … HBD1) -D1 #B2 #HB12 #HBD2 destruct
+ @(ex2_intro … (@B2.A)) // /2 width=1/
| #p #D1 #C1 #C2 #_ #IHC12 #d #M1 #H
- elim (lift_inv_appl … H) -H #B #A1 #H1 #H2 #H destruct
- elim (IHC12 ???) -IHC12 [4: // |2,3: skip ] #A2 #HA12 #H destruct (**) (* simplify line *)
- @(ex2_1_intro … (@B.A2)) // /2 width=1/
+ elim (lift_inv_appl … H) -H #B #A1 #H1 #HAC1 #H2
+ elim (IHC12 … HAC1) -C1 #A2 #HA12 #HAC2 destruct
+ @(ex2_intro … (@B.A2)) // /2 width=1/
]
qed-.
>(lift_inv_vref_lt … Hid2 … H) -M2 /3 width=3/
| elim (lift_inv_vref_ge … H) -H -Hd21 // -Hid2 #Hdh2i #H destruct
elim (le_inv_plus_l … Hdh2i) -Hdh2i #Hd2i #Hh2i
- @(ex2_1_intro … (#(i-h2))) [ /4 width=1/ ] -Hid1
+ @(ex2_intro … (#(i-h2))) [ /4 width=1/ ] -Hid1
>lift_vref_ge // -Hd2i /3 width=1/ (**) (* auto: needs some help here *)
]
| elim (le_inv_plus_l … Hid1) #Hd1i #Hh2i
elim (le_inv_plus_l … Hdh2i) #Hd2i #_
>(lift_vref_ge … Hid1) #H -Hid1
>(lift_inv_vref_ge_plus … H) -H /2 width=3/ -Hdh2i
- @(ex2_1_intro … (#(i-h2))) (**) (* auto: needs some help here *)
+ @(ex2_intro … (#(i-h2))) (**) (* auto: needs some help here *)
[ >lift_vref_ge // -Hd1i /3 width=1/
| >lift_vref_ge // -Hd2i -Hd1i /3 width=1/
]
| normalize #A1 #IHA1 #M2 #d1 #d2 #Hd21 #H
elim (lift_inv_abst … H) -H >plus_plus_comm_23 #A2 #HA12 #H destruct
elim (IHA1 … HA12) -IHA1 -HA12 /2 width=1/ -Hd21 #A #HA2 #HA1
- @(ex2_1_intro … (𝛌.A)) normalize //
+ @(ex2_intro … (𝛌.A)) normalize //
| normalize #B1 #A1 #IHB1 #IHA1 #M2 #d1 #d2 #Hd21 #H
elim (lift_inv_appl … H) -H #B2 #A2 #HB12 #HA12 #H destruct
elim (IHB1 … HB12) -IHB1 -HB12 // #B #HB2 #HB1
elim (IHA1 … HA12) -IHA1 -HA12 // -Hd21 #A #HA2 #HA1
- @(ex2_1_intro … (@B.A)) normalize //
+ @(ex2_intro … (@B.A)) normalize //
]
qed-.
definition liftable: predicate (relation term) ≝ λR.
∀h,M1,M2. R M1 M2 → ∀d. R (↑[d, h] M1) (↑[d, h] M2).
-definition deliftable: predicate (relation term) ≝ λR.
- ∀h,N1,N2. R N1 N2 → ∀d,M1. ↑[d, h] M1 = N1 →
- ∃∃M2. R M1 M2 & ↑[d, h] M2 = N2.
+definition deliftable_sn: predicate (relation term) ≝ λR.
+ ∀h,N1,N2. R N1 N2 → ∀d,M1. ↑[d, h] M1 = N1 →
+ ∃∃M2. R M1 M2 & ↑[d, h] M2 = N2.
+
+lemma star_liftable: ∀R. liftable R → liftable (star … R).
+#R #HR #h #M1 #M2 #H elim H -M2 // /3 width=3/
+qed.
+
+lemma star_deliftable_sn: ∀R. deliftable_sn R → deliftable_sn (star … R).
+#R #HR #h #N1 #N2 #H elim H -N2 /2 width=3/
+#N #N2 #_ #HN2 #IHN1 #d #M1 #HMN1
+elim (IHN1 … HMN1) -N1 #M #HM1 #HMN
+elim (HR … HN2 … HMN) -N /3 width=3/
+qed-.
+
+lemma lstar_liftable: ∀T,R. (∀t. liftable (R t)) →
+ ∀l. liftable (lstar T … R l).
+#T #R #HR #l #h #M1 #M2 #H elim H -l -M1 -M2 // /3 width=3/
+qed.
+
+lemma lstar_deliftable_sn: ∀T,R. (∀t. deliftable_sn (R t)) →
+ ∀l. deliftable_sn (lstar T … R l).
+#T #R #HR #l #h #N1 #N2 #H elim H -l -N1 -N2 /2 width=3/
+#t #N1 #N #HN1 #l #N2 #_ #IHN2 #d #M1 #HMN1
+elim (HR … HN1 … HMN1) -N1 #M #HM1 #HMN
+elim (IHN2 … HMN) -N /3 width=3/
+qed-.
#D #D #A #C #_ #_ #IHBD #IHAC #d <dsubst_lift_le // /2 width=1/
qed.
-lemma pred_inv_lift: deliftable pred.
+lemma pred_inv_lift: deliftable_sn pred.
#h #N1 #N2 #H elim H -N1 -N2 /2 width=3/
[ #C1 #C2 #_ #IHC12 #d #M1 #H
elim (lift_inv_abst … H) -H #A1 #HAC1 #H
elim (IHC12 … HAC1) -C1 #A2 #HA12 #HAC2 destruct
- @(ex2_1_intro … (𝛌.A2)) // /2 width=1/
+ @(ex2_intro … (𝛌.A2)) // /2 width=1/
| #D1 #D2 #C1 #C2 #_ #_ #IHD12 #IHC12 #d #M1 #H
elim (lift_inv_appl … H) -H #B1 #A1 #HBD1 #HAC1 #H
elim (IHD12 … HBD1) -D1 #B2 #HB12 #HBD2
elim (IHC12 … HAC1) -C1 #A2 #HA12 #HAC2 destruct
- @(ex2_1_intro … (@B2.A2)) // /2 width=1/
+ @(ex2_intro … (@B2.A2)) // /2 width=1/
| #D1 #D2 #C1 #C2 #_ #_ #IHD12 #IHC12 #d #M1 #H
elim (lift_inv_appl … H) -H #B1 #M #HBD1 #HM #H1
elim (lift_inv_abst … HM) -HM #A1 #HAC1 #H
elim (IHD12 … HBD1) -D1 #B2 #HB12 #HBD2
elim (IHC12 … HAC1) -C1 #A2 #HA12 #HAC2 destruct
- @(ex2_1_intro … ([⬐B2]A2)) /2 width=1/
+ @(ex2_intro … ([⬐B2]A2)) /2 width=1/
]
qed-.
| #C #B2 #C2 #HB2 #HC2 #H2 #HM2 #B1 #N #HB1 #H #HM1 destruct
@(pred_conf1_appl_beta … IH) // (**) (* /2 width=7 by pred_conf1_appl_beta/ does not work *)
| #B2 #N #B2 #H #HM2 #C #B1 #C1 #HB1 #HC1 #H1 #HM1 destruct
- @ex2_1_commute @(pred_conf1_appl_beta … IH) //
+ @ex2_commute @(pred_conf1_appl_beta … IH) //
| #C #B2 #C2 #HB2 #HC2 #H2 #HM2 #C0 #B1 #C1 #HB1 #HC1 #H1 #HM1 destruct
elim (IH B … HB1 … HB2) -HB1 -HB2 //
elim (IH C … HC1 … HC2) normalize // -B -C /3 width=5/
H : transient premise
IH : inductive premise
M, N : term
+R : arbitrary relation
+T, U : arbitrary small type
a, b : boolean
d, e : variable reference depth
h : relocation height
i, j : de Bruijn index
k : relocation height
+l : arbitrary list
m, n : measures on terms
p, q : redex pointer
r, s : redex pointer sequence
+t, u : arbitrary element
non associative with precedence 90
for @{'false}.
-lemma ex2_1_commute: ∀A0. ∀P0,P1:A0→Prop. (∃∃x0. P0 x0 & P1 x0) → ∃∃x0. P1 x0 & P0 x0.
-#A0 #P0 #P1 * /2 width=3/
-qed.
-
(* relations *)
definition confluent1: ∀A. relation A → predicate A ≝ λA,R,a0.
notation > "◊"
non associative with precedence 90
for @{'nil}.
-
-let rec Allr (A:Type[0]) (R:relation A) (l:list A) on l : Prop ≝
-match l with
-[ nil ⇒ True
-| cons a1 l ⇒ match l with [ nil ⇒ True | cons a2 _ ⇒ R a1 a2 ∧ Allr A R l ]
-].
| rpprec_nil : ∀b,q. rpprec (◊) (b::q)
| rppprc_cons: ∀p,q. rpprec (false::p) (true::q)
| rpprec_comp: ∀b,p,q. rpprec p q → rpprec (b::p) (b::q)
-| rpprec_skip: ∀p,q. rpprec (false::p) q → rpprec p q
+| rpprec_skip: rpprec (false::◊) ◊
.
interpretation "'precedes' on redex pointers"
interpretation "'less or equal to' on redex pointers"
'leq p q = (rple p q).
+
+lemma rpprec_rple: ∀p,q. p ≺ q → p ≤ q.
+/2 width=1/
+qed.
+
+lemma rple_false: false::◊ ≤ ◊.
+/2 width=1/
+qed.
+
+lemma rple_nil: ∀p. ◊ ≤ p.
+* // /2 width=1/
+qed.
+
+lemma rple_comp: ∀p1,p2. p1 ≤ p2 → ∀b. (b::p1) ≤ (b::p2).
+#p1 #p2 #H elim H -p2 // /3 width=3/
+qed.
(* Note: to us, a "normal" computation contracts redexes in non-decreasing positions *)
definition is_le: predicate rpseq ≝ λs.
Allr … rple s.
+
+(* Note: a normal spine computation *)
+definition is_spine_le: predicate rpseq ≝ λs.
+ is_spine s ∧ is_le s.
(* *)
(**************************************************************************)
-include "hap_computation.ma".
+include "labelled_hap_computation.ma".
(* KASHIMA'S "ST" COMPUTATION ***********************************************)
R. Kashima: "A proof of the Standization Theorem in λ-Calculus". Typescript note, (2000).
*)
inductive st: relation term ≝
-| st_vref: ∀M,i. hap M (#i) → st M (#i)
-| st_abst: ∀M,A,C. hap M (𝛌.A) → st A C → st M (𝛌.C)
-| st_appl: ∀M,B,D,A,C. hap M (@B.A) → st B D → st A C → st M (@D.C)
+| st_vref: ∀s,M,i. lhap s M (#i) → st M (#i)
+| st_abst: ∀s,M,A,C. lhap s M (𝛌.A) → st A C → st M (𝛌.C)
+| st_appl: ∀s,M,B,D,A,C. lhap s M (@B.A) → st B D → st A C → st M (@D.C)
.
notation "hvbox( @ term 46 C . break term 46 A )"
non associative with precedence 46
for @{ 'Application $C $A }.
+
+definition appl_compatible_dx: predicate (relation term) ≝ λR.
+ ∀B,A1,A2. R A1 A2 → R (@B.A1) (@B.A2).
+
+lemma star_appl_compatible_dx: ∀R. appl_compatible_dx R →
+ appl_compatible_dx (star … R).
+#R #HR #B #A1 #A2 #H elim H -A2 // /3 width=3/
+qed.
<key name="objects">xoa</key>
<key name="notations">xoa_notation</key>
<key name="include">basics/pts.ma</key>
- <key name="ex">2 1</key>
+ <key name="ex">3 1</key>
<key name="ex">3 2</key>
<key name="ex">4 3</key>
<key name="or">3</key>
include "basics/pts.ma".
-(* multiple existental quantifier (2, 1) *)
+(* multiple existental quantifier (3, 1) *)
-inductive ex2_1 (A0:Type[0]) (P0,P1:A0→Prop) : Prop ≝
- | ex2_1_intro: ∀x0. P0 x0 → P1 x0 → ex2_1 ? ? ?
+inductive ex3_1 (A0:Type[0]) (P0,P1,P2:A0→Prop) : Prop ≝
+ | ex3_1_intro: ∀x0. P0 x0 → P1 x0 → P2 x0 → ex3_1 ? ? ? ?
.
-interpretation "multiple existental quantifier (2, 1)" 'Ex P0 P1 = (ex2_1 ? P0 P1).
+interpretation "multiple existental quantifier (3, 1)" 'Ex P0 P1 P2 = (ex3_1 ? P0 P1 P2).
(* multiple existental quantifier (3, 2) *)
(* This file was generated by xoa.native: do not edit *********************)
-(* multiple existental quantifier (2, 1) *)
+(* multiple existental quantifier (3, 1) *)
-notation > "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1)"
+notation > "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1 break & term 19 P2)"
non associative with precedence 20
- for @{ 'Ex (λ${ident x0}.$P0) (λ${ident x0}.$P1) }.
+ for @{ 'Ex (λ${ident x0}.$P0) (λ${ident x0}.$P1) (λ${ident x0}.$P2) }.
-notation < "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1)"
+notation < "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1 break & term 19 P2)"
non associative with precedence 20
- for @{ 'Ex (λ${ident x0}:$T0.$P0) (λ${ident x0}:$T0.$P1) }.
+ for @{ 'Ex (λ${ident x0}:$T0.$P0) (λ${ident x0}:$T0.$P1) (λ${ident x0}:$T0.$P2) }.
(* multiple existental quantifier (3, 2) *)
@{ ${ fold right @{$Px} rec acc @{'exists (λ${ident x}.$acc)} } }
}.
+notation > "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1)"
+ non associative with precedence 20
+ for @{ 'exists2 (λ${ident x0}.$P0) (λ${ident x0}.$P1) }.
+
+notation < "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1)"
+ non associative with precedence 20
+ for @{ 'exists2 (λ${ident x0}:$T0.$P0) (λ${ident x0}:$T0.$P1) }.
+
(* sigma ********************************************************************)
notation < "hvbox(\Sigma ident i : ty break . p)"
]
] qed.
+let rec Allr (A:Type[0]) (R:relation A) (l:list A) on l : Prop ≝
+match l with
+[ nil ⇒ True
+| cons a1 l ⇒ match l with [ nil ⇒ True | cons a2 _ ⇒ R a1 a2 ∧ Allr A R l ]
+].
+
(**************************** Exists *******************************)
let rec Exists (A:Type[0]) (P:A → Prop) (l:list A) on l : Prop ≝
[ O ⇒ [ ]
| S m ⇒ a::(make_list A a m)
].
+
+(* ******** labelled reflexive and transitive closure ************)
+
+inductive lstar (A:Type[0]) (B:Type[0]) (R: A→relation B): list A → relation B ≝
+| lstar_nil : ∀b. lstar A B R ([]) b b
+| lstar_cons: ∀a,b1,b. R a b1 b →
+ ∀l,b2. lstar A B R l b b2 → lstar A B R (a::l) b1 b2
+.
+
+lemma lstar_step: ∀A,B,R,a,b1,b2. R a b1 b2 → lstar A B R ([a]) b1 b2.
+/2 width=3/
+qed.
+
+lemma lstar_inv_nil: ∀A,B,R,l,b1,b2. lstar A B R l b1 b2 → [] = l → b1 = b2.
+#A #B #R #l #b1 #b2 * -l -b1 -b2 //
+#a #b1 #b #_ #l #b2 #_ #H destruct
+qed-.
+
+lemma lstar_inv_cons: ∀A,B,R,l,b1,b2. lstar A B R l b1 b2 →
+ ∀a0,l0. a0::l0 = l →
+ ∃∃b. R a0 b1 b & lstar A B R l0 b b2.
+#A #B #R #l #b1 #b2 * -l -b1 -b2
+[ #b #a0 #l0 #H destruct
+| #a #b1 #b #Hb1 #l #b2 #Hb2 #a0 #l0 #H destruct /2 width=3/
+]
+qed-.
+
+lemma lstar_inv_step: ∀A,B,R,a,b1,b2. lstar A B R ([a]) b1 b2 → R a b1 b2.
+#A #B #R #a #b1 #b2 #H
+elim (lstar_inv_cons ?????? H ???) -H [4: // |2,3: skip ] #b #Hb1 #H (**) (* simplify line *)
+<(lstar_inv_nil ?????? H ?) -H // (**) (* simplify line *)
+qed-.
+
+theorem lstar_singlevalued: ∀A,B,R. (∀a. singlevalued ?? (R a)) →
+ ∀l. singlevalued … (lstar A B R l).
+#A #B #R #HR #l #b #c1 #H elim H -l -b -c1
+[ /2 width=5 by lstar_inv_nil/
+| #a #b #b1 #Hb1 #l #c1 #_ #IHbc1 #c2 #H
+ elim (lstar_inv_cons ?????? H ???) -H [4: // |2,3: skip ] #b2 #Hb2 #Hbc2 (**) (* simplify line *)
+ lapply (HR … Hb1 … Hb2) -b #H destruct /2 width=1/
+]
+qed-.
+
+theorem lstar_trans: ∀A,B,R,l1,b1,b. lstar A B R l1 b1 b →
+ ∀l2,b2. lstar A B R l2 b b2 → lstar A B R (l1@l2) b1 b2.
+#A #B #R #l1 #b1 #b #H elim H -l1 -b1 -b normalize // /3 width=3/
+qed-.
interpretation "exists" 'exists x = (ex ? x).
inductive ex2 (A:Type[0]) (P,Q:A →Prop) : Prop ≝
- ex_intro2: ∀ x:A. P x → Q x → ex2 A P Q.
+ ex2_intro: ∀ x:A. P x → Q x → ex2 A P Q.
+
+interpretation "exists on two predicates" 'exists2 x1 x2 = (ex2 ? x1 x2).
+
+lemma ex2_commute: ∀A0. ∀P0,P1:A0→Prop. (∃∃x0. P0 x0 & P1 x0) → ∃∃x0. P1 x0 & P0 x0.
+#A0 #P0 #P1 * /2 width=3/
+qed.
(* iff *)
definition iff :=
#a #b #c #Rab #sbc #sbc @(star_compl … Rab) //
qed.
-lemma star_ind_l :
- ∀A:Type[0].∀R:relation A.∀Q:A → A → Prop.
- (∀a.Q a a) →
- (∀a,b,c.R a b → star A R b c → Q b c → Q a c) →
- ∀a,b.star A R a b → Q a b.
-#A #R #Q #H1 #H2 #a #b #H0
-elim (star_to_starl ???? H0) // -H0 -b -a
-#a #b #c #Rab #slbc @H2 // @starl_to_star //
-qed.
+fact star_ind_l_aux: ∀A,R,a2. ∀P:predicate A.
+ P a2 →
+ (∀a1,a. R a1 a → star … R a a2 → P a → P a1) →
+ ∀a1,a. star … R a1 a → a = a2 → P a1.
+#A #R #a2 #P #H1 #H2 #a1 #a #Ha1
+elim (star_to_starl ???? Ha1) -a1 -a
+[ #a #b #c #Hab #Hbc #IH #H destruct /3 width=4/
+| #a #H destruct /2 width=1/
+]
+qed-.
+
+(* imporeved version of star_ind_l with "left_parameter" *)
+lemma star_ind_l: ∀A,R,a2. ∀P:predicate A.
+ P a2 →
+ (∀a1,a. R a1 a → star … R a a2 → P a → P a1) →
+ ∀a1. star … R a1 a2 → P a1.
+#A #R #a2 #P #H1 #H2 #a1 #Ha12
+@(star_ind_l_aux … H1 H2 … Ha12) //
+qed.
(* RC and star *)
include "turing/turing.ma".
-
+(* this no longer works: TODO
(* time *)
let rec time_of sig (M:TM sig) a b (p:computation sig M a b) on p ≝
match p with
∀l.∃c.∃p:computation sig M (init sig M l) c.
time_of … p ≤ f (|l|) ∧
(stop sig M c = true) ∧ (isnilb ? (out ?? c) = false).
-
+*)
(* Le macchine di Turing sono a Type[1], che vorrebbe un nuovo
esistenziale.
definition DTIME ≝ λsig:FinSet. λL: list sig → bool. λf.
∃M:TM sig. ComputeB sig M L ∧ ∃c. Time_Bound sig M (λx.c + c*(f x)).
*)
-
+(* this no longer works: TODO
inductive DTIME (sig:FinSet) (A:list sig → bool) (f:nat→nat) : Type[1] ≝
| DTIME_intro: ∀M:TM sig.∀c.
ComputeB_in_time sig M A (λx.c + c*(f x)) → DTIME sig A f.
-
+*)
(* space complexity *)
definition size_of_tape ≝ λsig.λtp: tape sig.|left sig tp|+|right sig tp|.
definition size_of_tapes ≝ λsig.λn.λtps: Vector (tape sig) n.
foldr ?? (λtp.λacc.max (size_of_tape sig tp) acc) 0 tps.
-
+(* this no longer works: TODO
definition size_of_config ≝ λsig.λM.λc.
size_of_tapes sig (S (tapes_no sig M)) (tapes sig M c).
∀l.∃c.∃p:computation sig M (init sig M l) c.
space_of … p ≤ f (|l|) ∧
(stop sig M c = true) ∧ (isnilb ? (out ?? c) = false).
-
+
inductive DSPACE (sig:FinSet) (A:list sig → bool) (f:nat→nat) : Type[1] ≝
| DTIME_intro: ∀M:TM sig.∀c.
ComputeB_in_space sig M A (λx.c + c*(f x)) → DSPACE sig A f.
-
+*)
\ No newline at end of file
#alpha #sep #inc #i #outc #Hloop
lapply (sem_while … (sem_mcr_step alpha sep) inc i outc Hloop) [%]
-Hloop * #t1 * #Hstar @(star_ind_l ??????? Hstar)
-[ #tapea whd in ⊢ (% → ?); #H1 #b #a #ls #rs #Htapea
+[ whd in ⊢ (% → ?); #H1 #b #a #ls #rs #Htapea
%
[ #Hb >Htapea in H1; >Hb #H1 cases (H1 ??)
[#_ #H2 >H2 % |*: % #H2 normalize in H2; destruct (H2)]
[#Hfalse @False_ind @(absurd ?? Hb) normalize in Hfalse; destruct %
|*:% #H2 normalize in H2; destruct (H2) ]
]
-| #tapea #tapeb #tapec #Hstar1 #HRtrue #IH #HRfalse
+| #tapeb #tapec #Hstar1 #HRtrue #IH #HRfalse
lapply (IH HRfalse) -IH -HRfalse whd in ⊢ (%→%); #IH
#b0 #a0 #ls #rs #Htapea cases Hstar1 #b * #bsep *
[* #a * #ls1 * #rs1 * >Htapea #H destruct (H) #Htapeb %
#alpha #sep #inc #i #outc #Hloop
lapply (sem_while … (sem_mcl_step alpha sep) inc i outc Hloop) [%]
-Hloop * #t1 * #Hstar @(star_ind_l ??????? Hstar)
-[ #tapea whd in ⊢ (% → ?); #H1 #b #a #ls #rs #Htapea
+[ whd in ⊢ (% → ?); #H1 #b #a #ls #rs #Htapea
%
[ #Hb >Htapea in H1; >Hb #H1 cases (H1 ??)
[#_ #H2 >H2 % |*: % #H2 normalize in H2; destruct (H2)]
[#Hfalse @False_ind @(absurd ?? Hb) normalize in Hfalse; destruct %
|*:% #H2 normalize in H2; destruct (H2) ]
]
-| #tapea #tapeb #tapec #Hstar1 #HRtrue #IH #HRfalse
+| #tapeb #tapec #Hstar1 #HRtrue #IH #HRfalse
lapply (IH HRfalse) -IH -HRfalse whd in ⊢ (%→%); #IH
#b0 #a0 #ls #rs #Htapea cases Hstar1 #b * #bsep *
[* #a * #ls1 * #rs1 * >Htapea #H destruct (H) #Htapeb %
#i #j #sig #n #is_endc #Hneq #Hi #Hj #ta #k #outc #Hloop
lapply (sem_while … (sem_comp_step i j sig n is_endc Hneq Hi Hj) … Hloop) //
-Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
-[ #tc whd in ⊢ (%→?); * * [ * [ *
+[ whd in ⊢ (%→?); * * [ * [ *
[* #curi * #Hcuri #Hendi #Houtc %
[ #_ @Houtc
| #ls #x #xs #ci #rs #ls0 #rs0 #Hnthi #Hnthj #Hnotendc
[ #_ @Houtc
| #ls #x #xs #ci #rs #ls0 #rs0 #_ #Hnthj >Hnthj in Hcj;
normalize in ⊢ (%→?); #H destruct (H) ] ]
- | #tc #td #te * #x * * * #Hendcx #Hci #Hcj #Hd #Hstar #IH #He lapply (IH He) -IH *
+ | #td #te * #x * * * #Hendcx #Hci #Hcj #Hd #Hstar #IH #He lapply (IH He) -IH *
#IH1 #IH2 %
[ >Hci >Hcj * [* #x0 * #H destruct (H) >Hendcx #H destruct (H)
|* [* #H @False_ind [cases H -H #H @H % | destruct (H)] | #H destruct (H)]]
copy src dst sig n is_sep ⊫ R_copy src dst sig n is_sep.
#src #dst #sig #n #is_sep #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
lapply (sem_while … (sem_copy_step src dst sig n is_sep Hneq Hsrc Hdst) … Hloop) //
--Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
-[ #tc whd in ⊢ (%→?); *
+-Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar -ta
+[ whd in ⊢ (%→?); *
[ * #x * * #Hx #Hsep #Houtc % [ %
[ #ls #x0 #xs #rs #sep #Hsrctc #Hnosep >Hsrctc in Hx; normalize in ⊢ (%→?);
#Hx0 destruct (Hx0) lapply (Hnosep ? (memb_hd …)) >Hsep
| #c #Hc #Hsepc @Houtc ]
| #_ @Houtc ]
]
-| #tc #td #te * #c0 * * #Hc0 #Hc0nosep #Hd #Hstar #IH #He
+| #td #te * #c0 * * #Hc0 #Hc0nosep #Hd #Hstar #IH #He
lapply (IH He) -IH * * #IH1 #IH2 #IH3 % [ %
[ #ls #x #xs #rs #sep #Hsrc_tc #Hnosep #Hsep #ls0 #x0 #target
#c #rs0 #Hlen #Hdst_tc
>Hsrc_tc in Hc0; normalize in ⊢ (%→?); #Hc0 destruct (Hc0)
- <(change_vec_same … tc src (niltape ?)) in Hd:(???(???(???%??)??));
- <(change_vec_same … tc dst (niltape ?)) in ⊢(???(???(???%??)??)→?);
- >Hdst_tc >Hsrc_tc >change_vec_change_vec >change_vec_change_vec
- >(change_vec_commute ?? tc ?? dst src) [|@(sym_not_eq … Hneq)]
+ <(change_vec_same … td src (niltape ?)) in Hd:(???(???(???%??)??));
+ <(change_vec_same … td dst (niltape ?)) in ⊢(???(???(???%??)??)→?);
+ >Hdst_tc >Hsrc_tc >(change_vec_change_vec ?) >change_vec_change_vec
+ >(change_vec_commute ?? td ?? dst src) [|@(sym_not_eq … Hneq)]
>change_vec_change_vec @(list_cases2 … Hlen)
[ #Hxsnil #Htargetnil #Hd>(IH2 … Hsep)
[ >Hd -Hd >Hxsnil >Htargetnil @(eq_vec … (niltape ?))
>nth_change_vec // >nth_change_vec // >Hxsnil % ]
|#hd1 #hd2 #tl1 #tl2 #Hxs #Htarget >Hxs >Htarget #Hd
>(IH1 (c0::ls) hd1 tl1 rs sep ?? Hsep (c0::ls0) hd2 tl2 c rs0)
- [ >Hd >(change_vec_commute … ?? tc ?? src dst) //
+ [ >Hd >(change_vec_commute … ?? td ?? src dst) //
>change_vec_change_vec
- >(change_vec_commute … ?? tc ?? dst src) [|@sym_not_eq //]
+ >(change_vec_commute … ?? td ?? dst src) [|@sym_not_eq //]
>change_vec_change_vec
>reverse_cons >associative_append >associative_append %
| >Hd >nth_change_vec // >nth_change_vec_neq // >Hdst_tc >Htarget //
#src #dst #sig #n #is_startc #is_endc #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
lapply (sem_while … (sem_match_step src dst sig n is_startc is_endc Hneq Hsrc Hdst) … Hloop) //
-Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
-[ #tc #Hfalse #ls #x #xs #end #rs #Hmid_src #Hnotend #Hend #Hnotstart
+[ #Hfalse #ls #x #xs #end #rs #Hmid_src #Hnotend #Hend #Hnotstart
cases (Hfalse … Hmid_src Hnotend Hend) -Hfalse
[(* current dest = None *) *
[ * #Hcur_dst #Houtc %
>reverse_cons >associative_append @(HFalse ?? Hnotnil)
]
]
-|-ta -tb #ta #tb #tc #Htrue #Hstar #IH #Hout lapply (IH Hout) -IH -Hout #IH whd
+|-ta #ta #tc #Htrue #Hstar #IH #Hout lapply (IH Hout) -IH -Hout #IH whd
#ls #x #xs #end #rs #Hmid_src #Hnotend #Hend #Hnotstart
lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?);
#src #dst #sig #n #is_sep #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
lapply (sem_while … (sem_parmove_step src dst sig n L is_sep Hneq Hsrc Hdst) … Hloop) //
-Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
-[ #tc whd in ⊢ (%→?); * #H #Houtc % [2: #_ @Houtc ] cases H
+[ whd in ⊢ (%→?); * #H #Houtc % [2: #_ @Houtc ] cases H
[ * [ * #x * #Hx #Hsep #ls #x0 #xs #rs #sep #Hsrctc #Hnosep >Hsrctc in Hx; normalize in ⊢ (%→?);
#Hx0 destruct (Hx0) lapply (Hnosep ? (memb_hd …)) >Hsep
#Hfalse destruct (Hfalse)
|#Hcur_dst #ls #x0 #xs #rs #sep #Hsrctc #Hnosep #Hsep #ls0 #x1 #target
#c #rs0 #Hlen #Hdsttc >Hdsttc in Hcur_dst; normalize in ⊢ (%→?); #H destruct (H)
]
-| #tc #td #te * #c0 * #c1 * * * #Hc0 #Hc1 #Hc0nosep #Hd #Hstar #IH #He
+| #td #te * #c0 * #c1 * * * #Hc0 #Hc1 #Hc0nosep #Hd #Hstar #IH #He
lapply (IH He) -IH * #IH1 #IH2 %
[ #ls #x #xs #rs #sep #Hsrc_tc #Hnosep #Hsep #ls0 #x0 #target
#c #rs0 #Hlen #Hdst_tc
]
| #hd1 #hd2 #tl1 #tl2 #Hxs #Htarget >Hxs >Htarget #Hd
>(IH1 ls hd1 tl1 (c0::rs) sep ?? Hsep ls0 hd2 tl2 c (x0::rs0))
- [ >Hd >(change_vec_commute … ?? tc ?? src dst) //
+ [ >Hd >(change_vec_commute … ?? td ?? src dst) //
>change_vec_change_vec
- >(change_vec_commute … ?? tc ?? dst src) [|@sym_not_eq //]
+ >(change_vec_commute … ?? td ?? dst src) [|@sym_not_eq //]
>change_vec_change_vec
>reverse_cons >associative_append
>reverse_cons >associative_append %
[ nil ⇒ False
| cons hd tl ⇒ a=hd ∨ mem A a tl
].
-
+(* this no longer works: TODO
definition reach ≝ λsig.λM:NTM sig.λc,c1:config sig M.
∃q,l,q1,mvs.
state ?? c = q ∧
mem ? 〈〈q,l〉,〈q1,mvs〉〉 (trans ? M) ∧
state ?? c1 = q1 ∧
tapes ?? c1 = (compose_vec ??? (tape_move sig) ? (tapes ?? c) mvs).
-
+*)
(*
definition empty_tapes ≝ λsig.λn.
mk_Vector ? n (make_list (tape sig) (mk_tape sig [] []) n) ?.
elim n // normalize //
qed.
*)
-
+(* this no longer works: TODO
definition init ≝ λsig.λM:NTM sig.λi:(list sig).
mk_config ??
(start sig M)
definition accepted ≝ λsig.λM:NTM sig.λw:(list sig).
∃c. star ? (reach sig M) (init sig M w) c ∧
accept ?? (state ?? c) = true.
-
+*)
\ No newline at end of file
[nil ⇒ None ?
|cons a _ ⇒ Some ? a
].
-
+(* this no longer works: TODO
definition tape_move ≝ λsig.λt: tape sig.λm:sig × move.
match \snd m with
[ R ⇒ mk_tape sig ((\fst m)::(left ? t)) (tail ? (right ? t))
| L ⇒ mk_tape sig (tail ? (left ? t)) ((\fst m)::(right ? t))
| N ⇒ mk_tape sig (left ? t) ((\fst m)::(tail ? (right ? t)))
].
-
+*)
definition current_chars ≝ λsig.λM:TM sig.λc:config sig M.
vec_map ?? (λt.option_hd ? (right ? t)) (S (tapes_no sig M)) (tapes ?? c).
[ None ⇒ l
| Some a ⇒ a::l
].
-
+(* this no longer works: TODO
definition step ≝ λsig.λM:TM sig.λc:config sig M.
let 〈news,mvs,outchar〉 ≝ trans sig M 〈state ?? c,current_chars ?? c〉 in
mk_config ??
(start sig M)
(vec_cons ? (mk_tape sig [] i) ? (empty_tapes sig (tapes_no sig M)))
[ ].
-
+*)
definition stop ≝ λsig.λM:TM sig.λc:config sig M.
halt sig M (state sig M c).
[ O ⇒ None ?
| S m ⇒ if p a then (Some ? a) else loop A m f p (f a)
].
-
+(* this no longer works: TODO
(* Compute ? M f states that f is computed by M *)
definition Compute ≝ λsig.λM:TM sig.λf:(list sig) → (list sig).
∀l.∃i.∃c.
∀l.∃i.∃c.
loop ? i (step sig M) (stop sig M) (init sig M l) = Some ? c ∧
(isnilb ? (out ?? c) = false).
+*)
\ No newline at end of file
qed.
lemma length_map: ∀A,B,l.∀f:A→B. length ? (map ?? f l) = length ? l.
-#A #B #l #f elim l // #a #tl #Hind normalize //
+#A #B #l #f elim l //
qed.
definition vec_map ≝ λA,B.λf:A→B.λn.λv:Vector A n.
STATO 9:
stato finale
*)
+(* TODO
match s with
[ q0 ⇒ match c with
[ bit x ⇒ 〈q1,〈mark x,L〉〉
| mark D ⇒ 〈q9,〈bit D,L〉〉 ]
| q9 ⇒ ? (* successo *)
].
-
+*)
(*
==================================
MACCHINE PER SPOSTARE LA "TESTINA"
STATO 4:
stato finale
*)
-
+(* TODO
match s with
[ q0 ⇒ match c with
[ bit x ⇒ 〈q1 x,〈c,L〉〉
| q2 d ⇒ 〈q0,〈d,L〉〉
| q3 ⇒ 〈q4,〈grid,L〉〉
| q4 ⇒ (* finale *) ].
-
+*)
(*
MACCHINA B
----------
STATO 4:
stato finale
*)
-
+(* TODO
match s with
[ q0 ⇒ match c with
[ bit x ⇒ 〈q1 x,〈c,R〉〉
| q2 d ⇒ 〈q0,〈d,R〉〉
| q3 ⇒ 〈q4,〈grid,L〉〉
| q4 ⇒ (* finale *) ].
-
+*)
(*
MACCHINA C
----------
stato finale
*)
-
+(* TODO
match s with
[ q0 ⇒ match c with
[ bit x ⇒ 〈q1 x,〈c,R〉〉
| _ ⇒ 〈q2 c,〈bit x,L〉〉
| q2 d ⇒ 〈q0,〈c,R〉〉
| q3 ⇒ (* finale *) ].
+*)
\ No newline at end of file
#intape #k #outc #Hloop
lapply (sem_while … sem_copy_step intape k outc Hloop) [%] -Hloop
* #ta * #Hstar @(star_ind_l ??????? Hstar)
-[ #tb whd in ⊢ (%→?); #Hleft
+[ whd in ⊢ (%→?); #Hleft
#ls #c #c0 #rs #l1 #l3 #l4 #Htb #Hl1nomarks #Hl3l4nomarks #Hlen #l1' #bv
#Hl1 #Hl1bits #l4' #bg #Hl4 #Hl4bits
cases (Hleft … Htb) -Hleft #Hc #Houtc % %
#Hl1bits lapply (Hl1bits 〈c',false〉 ?) [ @memb_hd ]
>Hc #Hfalse destruct ]
| @Houtc ]
-| #tb #tc #td whd in ⊢ (%→?→(?→%)→%→?); #Htc #Hstar1 #Hind #Htd
+| #tc #td whd in ⊢ (%→?→(?→%)→%→?); #Htc #Hstar1 #Hind #Htd
lapply (Hind Htd) -Hind #Hind
#ls #c #c0 #rs #l1 #l3 #l4 #Htb #Hl1nomarks #Hl3l4nomarks #Hlen #l1' #bv
#Hl1 #Hl1bits #l4' #bg #Hl4 #Hl4bits %2
#alpha #test #t #i #outc #Hloop
lapply (sem_while … (sem_atmr_step alpha test) t i outc Hloop) [%]
-Hloop * #t1 * #Hstar @(star_ind_l ??????? Hstar)
-[ #tapea * #Htapea *
+[ * #Htapea *
[ #H1 %
[#_ @Htapea
|#ls #c #rs #H2 >H2 in H1; whd in ⊢ (??%? → ?);
<Htapea //
]
]
-| #tapea #tapeb #tapec #Hleft #Hright #IH #HRfalse
+| #tapeb #tapec #Hleft #Hright #IH #HRfalse
lapply (IH HRfalse) -IH #IH %
[cases Hleft #ls * #a * #rs * * #Htapea #_ #_ >Htapea
whd in ⊢((??%?)→?); #H destruct (H);
#alpha #test #t #i #outc #Hloop
lapply (sem_while … (sem_atml_step alpha test) t i outc Hloop) [%]
-Hloop * #t1 * #Hstar @(star_ind_l ??????? Hstar)
-[ #tapea * #Htapea *
+[ * #Htapea *
[ #H1 %
[#_ @Htapea
|#ls #c #rs #H2 >H2 in H1; whd in ⊢ (??%? → ?);
]
]
]
-| #tapea #tapeb #tapec #Hleft #Hright #IH #HRfalse
+| #tapeb #tapec #Hleft #Hright #IH #HRfalse
lapply (IH HRfalse) -IH #IH %
[cases Hleft #ls0 * #a0 * #rs0 * * #Htapea #_ #_ >Htapea
whd in ⊢ ((??%?)→?); #H destruct (H)
#t #i #outc #Hloop
lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%]
-Hloop * #t1 * #Hstar @(star_ind_l ??????? Hstar)
-[ #tapea whd in ⊢ (%→?); #Rfalse #ls #c #rs #Htapea %
+[ whd in ⊢ (%→?); #Rfalse #ls #c #rs #Htapea %
[ %
[ #c' #Hc' #Hc lapply (Rfalse … Htapea) -Rfalse * >Hc
whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse)
| #b #b0 #bs #b0s #l1 #l2 #Hlen #Hbs1 #Hb0s1 #Hbs2 #Hb0s2 #Hl1 #Hc
cases (Rfalse … Htapea) -Rfalse >Hc whd in ⊢ (??%?→?);#Hfalse destruct (Hfalse)
]
-| #tapea #tapeb #tapec #Hleft #Hright #IH #Htapec lapply (IH Htapec) -Htapec -IH #IH
+| #tapeb #tapec #Hleft #Hright #IH #Htapec lapply (IH Htapec) -Htapec -IH #IH
whd in Hleft; #ls #c #rs #Htapea cases Hleft -Hleft
#ls0 * #c' * #rs0 * >Htapea #Hdes destruct (Hdes) * *
cases (true_or_false (bit_or_null c')) #Hc'
#intape #k #outc #Hloop
lapply (sem_while … sem_match_tuple_step intape k outc Hloop) [%] -Hloop
* #ta * #Hstar @(star_ind_l ??????? Hstar)
-[ #tb whd in ⊢ (%→?); #Hleft
+[ whd in ⊢ (%→?); #Hleft
#ls #cur #rs #Htb cases (Hleft … Htb) #Hgrid #Houtc %
[ #_ @Houtc
| #c #l1 #c1 #l2 #l3 #ls0 #rs0 #n #Hls #Hcur #Hrs
]
| (* in the interesting case, we execute a true iteration, then we restart the
while cycle, finally we end with a false iteration *)
- #tb #tc #td whd in ⊢ (%→?); #Htc
+ #tc #td whd in ⊢ (%→?); #Htc
#Hstar1 #IH whd in ⊢ (%→?); #Hright lapply (IH Hright) -IH whd in ⊢ (%→?); #IH
#ls #cur #rs #Htb %
[ (* cur can't be true because we assume at least one iteration *)
| (* match failed and there is no next tuple: the next while cycle will just exit *)
* * #Hdiff #Hnobars generalize in match (refl ? tc);
cases tc in ⊢ (???% → %);
- [ #_ normalize in ⊢ (??%?→?); #Hfalse destruct (Hfalse)
- |2,3: #x #xs #_ normalize in ⊢ (??%?→?); #Hfalse destruct (Hfalse) ]
- #ls1 #cur1 #rs1 #Htc normalize in ⊢ (??%?→?); #Hcur1
+ [ #_ @daemon (* long normalize *) (*
+ normalize in ⊢ (??%?→?); #Hfalse destruct (Hfalse)
+ *)
+ |2,3: #x #xs #_ @daemon (* long normalize *) (*
+ normalize in ⊢ (??%?→?); #Hfalse destruct (Hfalse) *) ]
+ #ls1 #cur1 #rs1 #Htc @daemon (* long normalize *) (*
+ normalize in ⊢ (??%?→?); #Hcur1
cases (IH … Htc) -IH #IH #_ %2 %
[ destruct (Hcur1) >IH [ >Htc % | % ]
| #l4 #newc #mv0 #l5
(* no_bars except the first one, where the tuple does not match ⇒
no match *)
@daemon
- ]
+ ] *)
]
]
qed.
lemma sem_uni_step :
accGRealize ? uni_step us_acc Pre_uni_step
R_uni_step_true R_uni_step_false.
+@daemon (* this no longer works: TODO *) (*
@(acc_sem_if_app_guarded STape … (sem_test_char ? (λc:STape.\fst c == bit false))
? (test_char_inv …) (sem_nop …) …)
[| @(sem_seq_app_guarded … (Realize_to_GRealize … sem_init_match) ???)
#b #Hb >Hb in Ht1; * #Hc #Ht1 lapply (Hc ? (refl ??)) -Hc #Hb' %
// cases b in Hb'; normalize #H1 //
]
+*)
qed.
(*
lemma sem_uni_step1:
uni_step ⊨ [us_acc: low_step_R_true, low_step_R_false].
+@daemon (* this no longer works: TODO *) (*
@(acc_Realize_to_acc_Realize … sem_uni_step)
[@unistep_true_to_low_step | @unistep_false_to_low_step ]
+*)
qed.
definition universalTM ≝ whileTM ? uni_step us_acc.
theorem sem_universal: ∀M:normalTM. ∀qstart.
universalTM ⊫ (low_R M qstart (R_TM FinBool M qstart)).
+@daemon (* this no longer works: TODO *) (*
#M #q #intape #i #outc #Hloop
lapply (sem_while … sem_uni_step1 intape i outc Hloop)
[@daemon] -Hloop
|@Houtc
]
]
+*)
qed.
theorem sem_universal2: ∀M:normalTM. ∀R.
#sig #S1 #S2 #outc cases outc #s #t %
qed.
+axiom daemon :∀P:Prop.P.
+
theorem sem_seq: ∀sig,M1,M2,R1,R2.
Realize sig M1 R1 → Realize sig M2 R2 →
Realize sig (seq sig M1 M2) (R1 ∘ R2).
+@daemon (* this no longer works: TODO *) (*
#sig #M1 #M2 #R1 #R2 #HR1 #HR2 #t #i #outc #Hloop
cases (HR1 t) #k1 * #outc1 * #Hloop1 #HM1
cases (HR2 (ctape sig (states ? M1) outc1)) #k2 * #outc2 * #Hloop2 #HM2
| @(ex_intro … (ctape ? (FinSum (states ? M1) (states ? M2)) (lift_confL … outc1)))
% // >eq_ctape_lift_conf_L >eq_ctape_lift_conf_R //
]
+*)
qed.