+++ /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 "turing/multi_universal/moves.ma".
-include "turing/if_multi.ma".
-include "turing/inject.ma".
-include "turing/basic_machines.ma".
-
-definition compare_states ≝ initN 3.
-
-definition comp0 : compare_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
-definition comp1 : compare_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
-definition comp2 : compare_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
-
-(*
-
-0) (x,x) → (x,x)(R,R) → 1
- (x,y≠x) → None 2
-1) (_,_) → None 1
-2) (_,_) → None 2
-
-*)
-
-definition trans_compare_step ≝
- λi,j.λsig:FinSet.λn.
- λp:compare_states × (Vector (option sig) (S n)).
- let 〈q,a〉 ≝ p in
- match pi1 … q with
- [ O ⇒ match nth i ? a (None ?) with
- [ None ⇒ 〈comp2,null_action sig n〉
- | Some ai ⇒ match nth j ? a (None ?) with
- [ None ⇒ 〈comp2,null_action ? n〉
- | Some aj ⇒ if ai == aj
- then 〈comp1,change_vec ? (S n)
- (change_vec ? (S n) (null_action ? n) (〈None ?,R〉) i)
- (〈None ?,R〉) j〉
- else 〈comp2,null_action ? n〉 ]
- ]
- | S q ⇒ match q with
- [ O ⇒ (* 1 *) 〈comp1,null_action ? n〉
- | S _ ⇒ (* 2 *) 〈comp2,null_action ? n〉 ] ].
-
-definition compare_step ≝
- λi,j,sig,n.
- mk_mTM sig n compare_states (trans_compare_step i j sig n)
- comp0 (λq.q == comp1 ∨ q == comp2).
-
-definition R_comp_step_true ≝
- λi,j,sig,n.λint,outt: Vector (tape sig) (S n).
- ∃x.
- current ? (nth i ? int (niltape ?)) = Some ? x ∧
- current ? (nth j ? int (niltape ?)) = Some ? x ∧
- outt = change_vec ??
- (change_vec ?? int
- (tape_move_right ? (nth i ? int (niltape ?))) i)
- (tape_move_right ? (nth j ? int (niltape ?))) j.
-
-definition R_comp_step_false ≝
- λi,j:nat.λsig,n.λint,outt: Vector (tape sig) (S n).
- (current ? (nth i ? int (niltape ?)) ≠ current ? (nth j ? int (niltape ?)) ∨
- current ? (nth i ? int (niltape ?)) = None ? ∨
- current ? (nth j ? int (niltape ?)) = None ?) ∧ outt = int.
-
-lemma comp_q0_q2_null :
- ∀i,j,sig,n,v.i < S n → j < S n →
- (nth i ? (current_chars ?? v) (None ?) = None ? ∨
- nth j ? (current_chars ?? v) (None ?) = None ?) →
- step sig n (compare_step i j sig n) (mk_mconfig ??? comp0 v)
- = mk_mconfig ??? comp2 v.
-#i #j #sig #n #v #Hi #Hj
-whd in ⊢ (? → ??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (?→??%?);
-* #Hcurrent
-[ @eq_f2
- [ whd in ⊢ (??(???%)?); >Hcurrent %
- | whd in ⊢ (??(????(???%))?); >Hcurrent @tape_move_null_action ]
-| @eq_f2
- [ whd in ⊢ (??(???%)?); >Hcurrent cases (nth i ?? (None sig)) //
- | whd in ⊢ (??(????(???%))?); >Hcurrent
- cases (nth i ?? (None sig)) [|#x] @tape_move_null_action ] ]
-qed.
-
-lemma comp_q0_q2_neq :
- ∀i,j,sig,n,v.i < S n → j < S n →
- (nth i ? (current_chars ?? v) (None ?) ≠ nth j ? (current_chars ?? v) (None ?)) →
- step sig n (compare_step i j sig n) (mk_mconfig ??? comp0 v)
- = mk_mconfig ??? comp2 v.
-#i #j #sig #n #v #Hi #Hj lapply (refl ? (nth i ?(current_chars ?? v)(None ?)))
-cases (nth i ?? (None ?)) in ⊢ (???%→?);
-[ #Hnth #_ @comp_q0_q2_null // % //
-| #ai #Hai lapply (refl ? (nth j ?(current_chars ?? v)(None ?)))
- cases (nth j ?? (None ?)) in ⊢ (???%→?);
- [ #Hnth #_ @comp_q0_q2_null // %2 //
- | #aj #Haj * #Hneq
- whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?); @eq_f2
- [ whd in match (trans ????); >Hai >Haj
- whd in ⊢ (??(???%)?); cut ((ai==aj)=false)
- [>(\bf ?) /2 by not_to_not/ % #Haiaj @Hneq
- >Hai >Haj //
- | #Haiaj >Haiaj % ]
- | whd in match (trans ????); >Hai >Haj
- whd in ⊢ (??(????(???%))?); cut ((ai==aj)=false)
- [>(\bf ?) /2 by not_to_not/ % #Haiaj @Hneq
- >Hai >Haj //
- |#Hcut >Hcut @tape_move_null_action
- ]
- ]
- ]
-]
-qed.
-
-lemma comp_q0_q1 :
- ∀i,j,sig,n,v,a.i ≠ j → i < S n → j < S n →
- nth i ? (current_chars ?? v) (None ?) = Some ? a →
- nth j ? (current_chars ?? v) (None ?) = Some ? a →
- step sig n (compare_step i j sig n) (mk_mconfig ??? comp0 v) =
- mk_mconfig ??? comp1
- (change_vec ? (S n)
- (change_vec ?? v
- (tape_move_right ? (nth i ? v (niltape ?))) i)
- (tape_move_right ? (nth j ? v (niltape ?))) j).
-#i #j #sig #n #v #a #Heq #Hi #Hj #Ha1 #Ha2
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?); @eq_f2
-[ whd in match (trans ????);
- >Ha1 >Ha2 whd in ⊢ (??(???%)?); >(\b ?) //
-| whd in match (trans ????);
- >Ha1 >Ha2 whd in ⊢ (??(????(???%))?); >(\b ?) //
- change with (change_vec ?????) in ⊢ (??(????%)?);
- <(change_vec_same … v j (niltape ?)) in ⊢ (??%?);
- <(change_vec_same … v i (niltape ?)) in ⊢ (??%?);
- >tape_move_multi_def
- >pmap_change >pmap_change <tape_move_multi_def
- >tape_move_null_action
- @eq_f2 // >nth_change_vec_neq //
-]
-qed.
-
-lemma sem_comp_step :
- ∀i,j,sig,n.i ≠ j → i < S n → j < S n →
- compare_step i j sig n ⊨
- [ comp1: R_comp_step_true i j sig n,
- R_comp_step_false i j sig n ].
-#i #j #sig #n #Hneq #Hi #Hj #int
-lapply (refl ? (current ? (nth i ? int (niltape ?))))
-cases (current ? (nth i ? int (niltape ?))) in ⊢ (???%→?);
-[ #Hcuri %{2} %
- [| % [ %
- [ whd in ⊢ (??%?); >comp_q0_q2_null /2/
- | normalize in ⊢ (%→?); #H destruct (H) ]
- | #_ % // % %2 // ] ]
-| #a #Ha lapply (refl ? (current ? (nth j ? int (niltape ?))))
- cases (current ? (nth j ? int (niltape ?))) in ⊢ (???%→?);
- [ #Hcurj %{2} %
- [| % [ %
- [ whd in ⊢ (??%?); >comp_q0_q2_null /2/ %2
- | normalize in ⊢ (%→?); #H destruct (H) ]
- | #_ % // >Ha >Hcurj % % % #H destruct (H) ] ]
- | #b #Hb %{2} cases (true_or_false (a == b)) #Hab
- [ %
- [| % [ %
- [whd in ⊢ (??%?); >(comp_q0_q1 … a Hneq Hi Hj) //
- >(\P Hab) <Hb @sym_eq @nth_vec_map
- | #_ whd >(\P Hab) %{b} % // % // <(\P Hab) // ]
- | * #H @False_ind @H %
- ] ]
- | %
- [| % [ %
- [whd in ⊢ (??%?); >comp_q0_q2_neq //
- <(nth_vec_map ?? (current …) i ? int (niltape ?))
- <(nth_vec_map ?? (current …) j ? int (niltape ?)) >Ha >Hb
- @(not_to_not ??? (\Pf Hab)) #H destruct (H) %
- | normalize in ⊢ (%→?); #H destruct (H) ]
- | #_ % // % % >Ha >Hb @(not_to_not ??? (\Pf Hab)) #H destruct (H) % ] ]
- ]
- ]
-]
-qed.
-
-definition compare ≝ λi,j,sig,n.
- whileTM … (compare_step i j sig n) comp1.
-
-(* (∃rs'.rs = rs0@rs' ∧ current ? (nth j ? outt (niltape ?)) = None ?) ∨
- (∃rs0'.rs0 = rs@rs0' ∧
- outt = change_vec ??
- (change_vec ?? int
- (mk_tape sig (reverse sig rs@x::ls) (None sig) []) i)
- (mk_tape sig (reverse sig rs@x::ls0) (option_hd sig rs0')
- (tail sig rs0')) j) ∨
- (∃xs,ci,cj,rs',rs0'.ci ≠ cj ∧ rs = xs@ci::rs' ∧ rs0 = xs@cj::rs0' ∧
- outt = change_vec ??
- (change_vec ?? int (midtape sig (reverse ? xs@x::ls) ci rs') i)
- (midtape sig (reverse ? xs@x::ls0) cj rs0') j)).*)
-definition R_compare ≝
- λi,j,sig,n.λint,outt: Vector (tape sig) (S n).
- ((current ? (nth i ? int (niltape ?)) ≠ current ? (nth j ? int (niltape ?)) ∨
- current ? (nth i ? int (niltape ?)) = None ? ∨
- current ? (nth j ? int (niltape ?)) = None ?) → outt = int) ∧
- (∀ls,x,rs,ls0,rs0.
-(* nth i ? int (niltape ?) = midtape sig ls x (xs@ci::rs) → *)
- nth i ? int (niltape ?) = midtape sig ls x rs →
- nth j ? int (niltape ?) = midtape sig ls0 x rs0 →
- (∃rs'.rs = rs0@rs' ∧
- outt = change_vec ??
- (change_vec ?? int
- (mk_tape sig (reverse sig rs0@x::ls) (option_hd sig rs') (tail ? rs')) i)
- (mk_tape sig (reverse sig rs0@x::ls0) (None ?) [ ]) j) ∨
- (∃rs0'.rs0 = rs@rs0' ∧
- outt = change_vec ??
- (change_vec ?? int
- (mk_tape sig (reverse sig rs@x::ls) (None sig) []) i)
- (mk_tape sig (reverse sig rs@x::ls0) (option_hd sig rs0')
- (tail sig rs0')) j) ∨
- (∃xs,ci,cj,rs',rs0'.ci ≠ cj ∧ rs = xs@ci::rs' ∧ rs0 = xs@cj::rs0' ∧
- outt = change_vec ??
- (change_vec ?? int (midtape sig (reverse ? xs@x::ls) ci rs') i)
- (midtape sig (reverse ? xs@x::ls0) cj rs0') j)).
-
-lemma wsem_compare : ∀i,j,sig,n.i ≠ j → i < S n → j < S n →
- compare i j sig n ⊫ R_compare i j sig n.
-#i #j #sig #n #Hneq #Hi #Hj #ta #k #outc #Hloop
-lapply (sem_while … (sem_comp_step i j sig n Hneq Hi Hj) … Hloop) //
--Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
-[ whd in ⊢ (%→?); * * [ *
- [ #Hcicj #Houtc %
- [ #_ @Houtc
- | #ls #x #rs #ls0 #rs0 #Hnthi #Hnthj
- >Hnthi in Hcicj; >Hnthj normalize in ⊢ (%→?); * #H @False_ind @H %
- ]
- | #Hci #Houtc %
- [ #_ @Houtc
- | #ls #x #rs #ls0 #rs0 #Hnthi >Hnthi in Hci;
- normalize in ⊢ (%→?); #H destruct (H) ] ]
- | #Hcj #Houtc %
- [ #_ @Houtc
- | #ls #x #rs #ls0 #rs0 #_ #Hnthj >Hnthj in Hcj;
- normalize in ⊢ (%→?); #H destruct (H) ] ]
-| #td #te * #x * * #Hci #Hcj #Hd #Hstar #IH #He lapply (IH He) -IH *
- #IH1 #IH2 %
- [ >Hci >Hcj * [ *
- [ * #H @False_ind @H % | #H destruct (H)] | #H destruct (H)]
- | #ls #c0 #rs #ls0 #rs0 cases rs
- [ -IH2 #Hnthi #Hnthj % %2 %{rs0} % [%]
- >Hnthi in Hd; #Hd >Hd in IH1; #IH1 >IH1
- [| % %2 >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec // % ]
- >Hnthj cases rs0 [| #r1 #rs1 ] %
- | #r1 #rs1 #Hnthi cases rs0
- [ -IH2 #Hnthj % % %{(r1::rs1)} % [%]
- >Hnthj in Hd; #Hd >Hd in IH1; #IH1 >IH1
- [| %2 >nth_change_vec // ]
- >Hnthi >Hnthj %
- | #r2 #rs2 #Hnthj lapply IH2; >Hd in IH1; >Hnthi >Hnthj
- >nth_change_vec //
- >nth_change_vec_neq [| @sym_not_eq // ] >nth_change_vec //
- cases (true_or_false (r1 == r2)) #Hr1r2
- [ >(\P Hr1r2) #_ #IH2 cases (IH2 … (refl ??) (refl ??)) [ *
- [ * #rs' * #Hrs1 #Hcurout_j % % %{rs'}
- >Hrs1 %
- [ %
- | >Hcurout_j >change_vec_commute // >change_vec_change_vec
- >change_vec_commute // @sym_not_eq // ]
- | * #rs0' * #Hrs2 #Hcurout_i % %2 %{rs0'}
- >Hrs2 >Hcurout_i % //
- >change_vec_commute // >change_vec_change_vec
- >change_vec_commute [|@sym_not_eq//] >change_vec_change_vec
- >reverse_cons >associative_append >associative_append % ]
- | * #xs * #ci * #cj * #rs' * #rs0' * * * #Hcicj #Hrs1 #Hrs2
- >change_vec_commute // >change_vec_change_vec
- >change_vec_commute [| @sym_not_eq ] // >change_vec_change_vec
- #Houtc %2 %{(r2::xs)} %{ci} %{cj} %{rs'} %{rs0'}
- % [ % [ % [ // | >Hrs1 // ] | >Hrs2 // ]
- | >reverse_cons >associative_append >associative_append >Houtc % ] ]
- | lapply (\Pf Hr1r2) -Hr1r2 #Hr1r2 #IH1 #_ %2
- >IH1 [| % % normalize @(not_to_not … Hr1r2) #H destruct (H) % ]
- %{[]} %{r1} %{r2} %{rs1} %{rs2} % [ % [ % /2/ | % ] | % ] ]]]]]
-qed.
-
-lemma terminate_compare : ∀i,j,sig,n,t.
- i ≠ j → i < S n → j < S n →
- compare i j sig n ↓ t.
-#i #j #sig #n #t #Hneq #Hi #Hj
-@(terminate_while … (sem_comp_step …)) //
-<(change_vec_same … t i (niltape ?))
-cases (nth i (tape sig) t (niltape ?))
-[ % #t1 * #x * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct
-|2,3: #a0 #al0 % #t1 * #x * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct
-| #ls #c #rs lapply c -c lapply ls -ls lapply t -t elim rs
- [#t #ls #c % #t1 * #x * * >nth_change_vec // normalize in ⊢ (%→?);
- #H1 destruct (H1) #_ >change_vec_change_vec #Ht1 %
- #t2 * #x0 * * >Ht1 >nth_change_vec_neq [|@sym_not_eq //]
- >nth_change_vec // normalize in ⊢ (%→?); #H destruct (H)
- |#r0 #rs0 #IH #t #ls #c % #t1 * #x * * >nth_change_vec //
- normalize in ⊢ (%→?); #H destruct (H) #Hcur
- >change_vec_change_vec >change_vec_commute // #Ht1 >Ht1 @IH
- ]
-]
-qed.
-
-lemma sem_compare : ∀i,j,sig,n.
- i ≠ j → i < S n → j < S n →
- compare i j sig n ⊨ R_compare i j sig n.
-#i #j #sig #n #Hneq #Hi #Hj @WRealize_to_Realize
- [/2/| @wsem_compare // ]
-qed.
+++ /dev/null
-(*
- ||M|| This file is part of HELM, an Hypertextual, Electronic
- ||A|| Library of Mathematics, developed at the Computer Science
- ||T|| Department of the University of Bologna, Italy.
- ||I||
- ||T||
- ||A||
- \ / This file is distributed under the terms of the
- \ / GNU General Public License Version 2
- V_____________________________________________________________*)
-
-include "turing/multi_universal/moves.ma".
-include "turing/if_multi.ma".
-include "turing/inject.ma".
-include "turing/basic_machines.ma".
-
-definition copy_states ≝ initN 3.
-
-definition copy0 : copy_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
-definition copy1 : copy_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
-definition copy2 : copy_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
-
-
-definition trans_copy_step ≝
- λsrc,dst.λsig:FinSet.λn.
- λp:copy_states × (Vector (option sig) (S n)).
- let 〈q,a〉 ≝ p in
- match pi1 … q with
- [ O ⇒ match nth src ? a (None ?) with
- [ None ⇒ 〈copy2,null_action sig n〉
- | Some ai ⇒ match nth dst ? a (None ?) with
- [ None ⇒ 〈copy2,null_action ? n〉
- | Some aj ⇒
- 〈copy1,change_vec ? (S n)
- (change_vec ? (S n) (null_action ? n) (〈None ?,R〉) src)
- (〈Some ? ai,R〉) dst〉
- ]
- ]
- | S q ⇒ match q with
- [ O ⇒ (* 1 *) 〈copy1,null_action ? n〉
- | S _ ⇒ (* 2 *) 〈copy2,null_action ? n〉 ] ].
-
-definition copy_step ≝
- λsrc,dst,sig,n.
- mk_mTM sig n copy_states (trans_copy_step src dst sig n)
- copy0 (λq.q == copy1 ∨ q == copy2).
-
-definition R_copy_step_true ≝
- λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
- ∃x,y.
- current ? (nth src ? int (niltape ?)) = Some ? x ∧
- current ? (nth dst ? int (niltape ?)) = Some ? y ∧
- outt = change_vec ??
- (change_vec ?? int
- (tape_move_mono ? (nth src ? int (niltape ?)) 〈None ?, R〉) src)
- (tape_move_mono ? (nth dst ? int (niltape ?)) 〈Some ? x, R〉) dst.
-
-definition R_copy_step_false ≝
- λsrc,dst:nat.λsig,n.λint,outt: Vector (tape sig) (S n).
- (current ? (nth src ? int (niltape ?)) = None ? ∨
- current ? (nth dst ? int (niltape ?)) = None ?) ∧ outt = int.
-
-lemma copy_q0_q2_null :
- ∀src,dst,sig,n,v.src < S n → dst < S n →
- (nth src ? (current_chars ?? v) (None ?) = None ? ∨
- nth dst ? (current_chars ?? v) (None ?) = None ?) →
- step sig n (copy_step src dst sig n) (mk_mconfig ??? copy0 v)
- = mk_mconfig ??? copy2 v.
-#src #dst #sig #n #v #Hi #Hj
-whd in ⊢ (? → ??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (?→??%?);
-* #Hcurrent
-[ @eq_f2
- [ whd in ⊢ (??(???%)?); >Hcurrent %
- | whd in ⊢ (??(????(???%))?); >Hcurrent @tape_move_null_action ]
-| @eq_f2
- [ whd in ⊢ (??(???%)?); >Hcurrent cases (nth src ?? (None sig)) //
- | whd in ⊢ (??(????(???%))?); >Hcurrent
- cases (nth src ?? (None sig)) [|#x] @tape_move_null_action ] ]
-qed.
-
-lemma copy_q0_q1 :
- ∀src,dst,sig,n,v,a,b.src ≠ dst → src < S n → dst < S n →
- nth src ? (current_chars ?? v) (None ?) = Some ? a →
- nth dst ? (current_chars ?? v) (None ?) = Some ? b →
- step sig n (copy_step src dst sig n) (mk_mconfig ??? copy0 v) =
- mk_mconfig ??? copy1
- (change_vec ? (S n)
- (change_vec ?? v
- (tape_move_mono ? (nth src ? v (niltape ?)) 〈None ?, R〉) src)
- (tape_move_mono ? (nth dst ? v (niltape ?)) 〈Some ? a, R〉) dst).
-#src #dst #sig #n #v #a #b #Heq #Hsrc #Hdst #Ha1 #Ha2
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?); @eq_f2
-[ whd in match (trans ????);
- >Ha1 >Ha2 whd in ⊢ (??(???%)?); >(\b ?) //
-| whd in match (trans ????);
- >Ha1 >Ha2 whd in ⊢ (??(????(???%))?); >(\b ?) //
- change with (change_vec ?????) in ⊢ (??(????%)?);
- <(change_vec_same … v dst (niltape ?)) in ⊢ (??%?);
- <(change_vec_same … v src (niltape ?)) in ⊢ (??%?);
- >tape_move_multi_def
- >pmap_change >pmap_change <tape_move_multi_def
- >tape_move_null_action
- @eq_f2 // >nth_change_vec_neq //
-]
-qed.
-
-lemma sem_copy_step :
- ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n →
- copy_step src dst sig n ⊨
- [ copy1: R_copy_step_true src dst sig n,
- R_copy_step_false src dst sig n ].
-#src #dst #sig #n #Hneq #Hsrc #Hdst #int
-lapply (refl ? (current ? (nth src ? int (niltape ?))))
-cases (current ? (nth src ? int (niltape ?))) in ⊢ (???%→?);
-[ #Hcur_src %{2} %
- [| % [ %
- [ whd in ⊢ (??%?); >copy_q0_q2_null /2/
- | normalize in ⊢ (%→?); #H destruct (H) ]
- | #_ % // % // ] ]
-| #a #Ha lapply (refl ? (current ? (nth dst ? int (niltape ?))))
- cases (current ? (nth dst ? int (niltape ?))) in ⊢ (???%→?);
- [ #Hcur_dst %{2} %
- [| % [ %
- [ whd in ⊢ (??%?); >copy_q0_q2_null /2/
- | normalize in ⊢ (%→?); #H destruct (H) ]
- | #_ % // %2 >Hcur_dst % ] ]
- | #b #Hb %{2} %
- [| % [ %
- [whd in ⊢ (??%?); >(copy_q0_q1 … a b Hneq Hsrc Hdst) //
- | #_ %{a} %{b} % // % //]
- | * #H @False_ind @H %
- ]
- ]
- ]
-]
-qed.
-
-definition copy ≝ λsrc,dst,sig,n.
- whileTM … (copy_step src dst sig n) copy1.
-
-definition R_copy ≝
- λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
- ((current ? (nth src ? int (niltape ?)) = None ? ∨
- current ? (nth dst ? int (niltape ?)) = None ?) → outt = int) ∧
- (∀ls,x,x0,rs,ls0,rs0.
- nth src ? int (niltape ?) = midtape sig ls x rs →
- nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
- (∃rs01,rs02.rs0 = rs01@rs02 ∧ |rs01| = |rs| ∧
- outt = change_vec ??
- (change_vec ?? int
- (mk_tape sig (reverse sig rs@x::ls) (None sig) []) src)
- (mk_tape sig (reverse sig rs@x::ls0) (option_hd sig rs02)
- (tail sig rs02)) dst) ∨
- (∃rs1,rs2.rs = rs1@rs2 ∧ |rs1| = |rs0| ∧
- outt = change_vec ??
- (change_vec ?? int
- (mk_tape sig (reverse sig rs1@x::ls) (option_hd sig rs2)
- (tail sig rs2)) src)
- (mk_tape sig (reverse sig rs1@x::ls0) (None sig) []) dst)).
-
-lemma wsem_copy : ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n →
- copy src dst sig n ⊫ R_copy src dst sig n.
-#src #dst #sig #n #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
-lapply (sem_while … (sem_copy_step src dst sig n Hneq Hsrc Hdst) … Hloop) //
--Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
-[ whd in ⊢ (%→?); * #Hnone #Hout %
- [#_ @Hout
- |#ls #x #x0 #rs #ls0 #rs0 #Hsrc1 #Hdst1 @False_ind cases Hnone
- [>Hsrc1 normalize #H destruct (H) | >Hdst1 normalize #H destruct (H)]
- ]
-|#tc #td * #x * #y * * #Hcx #Hcy #Htd #Hstar #IH #He lapply (IH He) -IH *
- #IH1 #IH2 %
- [* [>Hcx #H destruct (H) | >Hcy #H destruct (H)]
- |#ls #x' #y' #rs #ls0 #rs0 #Hnth_src #Hnth_dst
- >Hnth_src in Hcx; whd in ⊢ (??%?→?); #H destruct (H)
- >Hnth_dst in Hcy; whd in ⊢ (??%?→?); #H destruct (H)
- >Hnth_src in Htd; >Hnth_dst -Hnth_src -Hnth_dst
- cases rs
- [(* the source tape is empty after the move *)
- #Htd lapply (IH1 ?)
- [%1 >Htd >nth_change_vec_neq [2:@(not_to_not … Hneq) //] >nth_change_vec //]
- #Hout (* whd in match (tape_move ???); *) %1 %{([])} %{rs0} %
- [% [// | // ]
- |whd in match (reverse ??); whd in match (reverse ??);
- >Hout >Htd @eq_f2 // cases rs0 //
- ]
- |#c1 #tl1 cases rs0
- [(* the dst tape is empty after the move *)
- #Htd lapply (IH1 ?) [%2 >Htd >nth_change_vec //]
- #Hout (* whd in match (tape_move ???); *) %2 %{[ ]} %{(c1::tl1)} %
- [% [// | // ]
- |whd in match (reverse ??); whd in match (reverse ??);
- >Hout >Htd @eq_f2 //
- ]
- |#c2 #tl2 whd in match (tape_move_mono ???); whd in match (tape_move_mono ???);
- #Htd
- cut (nth src (tape sig) td (niltape sig)=midtape sig (x::ls) c1 tl1)
- [>Htd >nth_change_vec_neq [2:@(not_to_not … Hneq) //] @nth_change_vec //]
- #Hsrc_td
- cut (nth dst (tape sig) td (niltape sig)=midtape sig (x::ls0) c2 tl2)
- [>Htd @nth_change_vec //]
- #Hdst_td cases (IH2 … Hsrc_td Hdst_td) -Hsrc_td -Hdst_td
- [* #rs01 * #rs02 * * #H1 #H2 #H3 %1
- %{(c2::rs01)} %{rs02} % [% [@eq_f //|normalize @eq_f @H2]]
- >Htd in H3; >change_vec_commute // >change_vec_change_vec
- >change_vec_commute [2:@(not_to_not … Hneq) //] >change_vec_change_vec
- #H >reverse_cons >associative_append >associative_append @H
- |* #rs11 * #rs12 * * #H1 #H2 #H3 %2
- %{(c1::rs11)} %{rs12} % [% [@eq_f //|normalize @eq_f @H2]]
- >Htd in H3; >change_vec_commute // >change_vec_change_vec
- >change_vec_commute [2:@(not_to_not … Hneq) //] >change_vec_change_vec
- #H >reverse_cons >associative_append >associative_append @H
- ]
- ]
- ]
- ]
-qed.
-
-
-lemma terminate_copy : ∀src,dst,sig,n,t.
- src ≠ dst → src < S n → dst < S n → copy src dst sig n ↓ t.
-#src #dst #sig #n #t #Hneq #Hsrc #Hdts
-@(terminate_while … (sem_copy_step …)) //
-<(change_vec_same … t src (niltape ?))
-cases (nth src (tape sig) t (niltape ?))
-[ % #t1 * #x * #y * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct
-|2,3: #a0 #al0 % #t1 * #x * #y * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct
-| #ls #c #rs lapply c -c lapply ls -ls lapply t -t elim rs
- [#t #ls #c % #t1 * #x * #y * * >nth_change_vec // normalize in ⊢ (%→?);
- #H1 destruct (H1) #_ >change_vec_change_vec #Ht1 %
- #t2 * #x0 * #y0 * * >Ht1 >nth_change_vec_neq [|@sym_not_eq //]
- >nth_change_vec // normalize in ⊢ (%→?); #H destruct (H)
- |#r0 #rs0 #IH #t #ls #c % #t1 * #x * #y * * >nth_change_vec //
- normalize in ⊢ (%→?); #H destruct (H) #Hcur
- >change_vec_change_vec >change_vec_commute // #Ht1 >Ht1 @IH
- ]
-]
-qed.
-
-lemma sem_copy : ∀src,dst,sig,n.
- src ≠ dst → src < S n → dst < S n →
- copy src dst sig n ⊨ R_copy src dst sig n.
-#i #j #sig #n #Hneq #Hi #Hj @WRealize_to_Realize [/2/| @wsem_copy // ]
-qed.
(* *)
(**************************************************************************)
-include "turing/simple_machines.ma".
-include "turing/multi_universal/compare.ma".
-include "turing/multi_universal/par_test.ma".
-include "turing/multi_universal/moves_2.ma".
+include "turing/auxiliary_multi_machines.ma".
-lemma eq_vec_change_vec : ∀sig,n.∀v1,v2:Vector sig n.∀i,t,d.
- nth i ? v2 d = t →
- (∀j.i ≠ j → nth j ? v1 d = nth j ? v2 d) →
- v2 = change_vec ?? v1 t i.
-#sig #n #v1 #v2 #i #t #d #H1 #H2 @(eq_vec … d)
-#i0 #Hlt cases (decidable_eq_nat i0 i) #Hii0
-[ >Hii0 >nth_change_vec //
-| >nth_change_vec_neq [|@sym_not_eq //] @sym_eq @H2 @sym_not_eq // ]
-qed.
-
-lemma right_mk_tape :
- ∀sig,ls,c,rs.(c = None ? → ls = [ ] ∨ rs = [ ]) → right ? (mk_tape sig ls c rs) = rs.
-#sig #ls #c #rs cases c // cases ls
-[ cases rs //
-| #l0 #ls0 #H normalize cases (H (refl ??)) #H1 [ destruct (H1) | >H1 % ] ]
-qed.
-
-lemma left_mk_tape : ∀sig,ls,c,rs.left ? (mk_tape sig ls c rs) = ls.
-#sig #ls #c #rs cases c // cases ls // cases rs //
-qed.
-
-lemma current_mk_tape : ∀sig,ls,c,rs.current ? (mk_tape sig ls c rs) = c.
-#sig #ls #c #rs cases c // cases ls // cases rs //
-qed.
-
-lemma length_tail : ∀A,l.0 < |l| → |tail A l| < |l|.
-#A * normalize //
-qed.
-
-(*
-[ ] → [ ], l2, 1
-a::al →
- [ ] → [ ], l1, 2
- b::bl → match rec(al,bl)
- x, y, 1 → b::x, y, 1
- x, y, 2 → a::x, y, 2
-*)
-
-lemma lists_length_split :
- ∀A.∀l1,l2:list A.(∃la,lb.(|la| = |l1| ∧ l2 = la@lb) ∨ (|la| = |l2| ∧ l1 = la@lb)).
-#A #l1 elim l1
-[ #l2 %{[ ]} %{l2} % % %
-| #hd1 #tl1 #IH *
- [ %{[ ]} %{(hd1::tl1)} %2 % %
- | #hd2 #tl2 cases (IH tl2) #x * #y *
- [ * #IH1 #IH2 %{(hd2::x)} %{y} % normalize % //
- | * #IH1 #IH2 %{(hd1::x)} %{y} %2 normalize % // ]
- ]
-]
-qed.
-
-definition option_cons ≝ λsig.λc:option sig.λl.
- match c with [ None ⇒ l | Some c0 ⇒ c0::l ].
-
-lemma opt_cons_tail_expand : ∀A,l.l = option_cons A (option_hd ? l) (tail ? l).
-#A * //
-qed.
-
-definition match_test ≝ λsrc,dst.λsig:DeqSet.λn.λv:Vector ? n.
- match (nth src (option sig) v (None ?)) with
- [ None ⇒ false
- | Some x ⇒ notb (nth dst (DeqOption sig) v (None ?) == None ?) ].
-
+(* rewind *)
definition rewind ≝ λsrc,dst,sig,n.
parmove src dst sig n L · mmove src sig n R · mmove dst sig n R.
∀ls0,y,rs0.nth dst ? int (niltape ?) = midtape sig ls0 y rs0 →
outt = int).
-(*
-theorem accRealize_to_Realize :
- ∀sig,n.∀M:mTM sig n.∀Rtrue,Rfalse,acc.
- M ⊨ [ acc: Rtrue, Rfalse ] → M ⊨ Rtrue ∪ Rfalse.
-#sig #n #M #Rtrue #Rfalse #acc #HR #t
-cases (HR t) #k * #outc * * #Hloop
-#Htrue #Hfalse %{k} %{outc} % //
-cases (true_or_false (cstate sig (states sig n M) n outc == acc)) #Hcase
-[ % @Htrue @(\P Hcase) | %2 @Hfalse @(\Pf Hcase) ]
-qed.
-*)
-
lemma sem_rewind_strong : ∀src,dst,sig,n.
src ≠ dst → src < S n → dst < S n →
rewind src dst sig n ⊨ R_rewind_strong src dst sig n.
#ta #tb * * * #H1 #H2 #H3 #H4 % /2 by /
qed.
+(* match step *)
+
+definition match_test ≝ λsrc,dst.λsig:DeqSet.λn.λv:Vector ? n.
+ match (nth src (option sig) v (None ?)) with
+ [ None ⇒ false
+ | Some x ⇒ notb (nth dst (DeqOption sig) v (None ?) == None ?) ].
+
definition match_step ≝ λsrc,dst,sig,n.
compare src dst sig n ·
(ifTM ?? (partest sig n (match_test src dst sig ?))
+++ /dev/null
-(*
- ||M|| This file is part of HELM, an Hypertextual, Electronic
- ||A|| Library of Mathematics, developed at the Computer Science
- ||T|| Department of the University of Bologna, Italy.
- ||I||
- ||T||
- ||A||
- \ / This file is distributed under the terms of the
- \ / GNU General Public License Version 2
- V_____________________________________________________________*)
-
-include "turing/turing.ma".
-include "turing/inject.ma".
-include "turing/while_multi.ma".
-
-definition parmove_states ≝ initN 3.
-
-definition parmove0 : parmove_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
-definition parmove1 : parmove_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
-definition parmove2 : parmove_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
-
-(*
-
-src: a b c ... z # ---→ a b c ... z #
- ^ ^
-
-dst: _ _ _ ... _ d ---→ a b c ... z d
- ^ ^
-
-0) (x ≠ sep,_) → (x,x)(R,R) → 1
- (sep,d) → None 2
-1) (_,_) → None 1
-2) (_,_) → None 2
-
-*)
-
-definition trans_parmove_step ≝
- λsrc,dst,sig,n,D,is_sep.
- λp:parmove_states × (Vector (option sig) (S n)).
- let 〈q,a〉 ≝ p in
- match pi1 … q with
- [ O ⇒ match nth src ? a (None ?) with
- [ None ⇒ 〈parmove2,null_action ? n〉
- | Some a0 ⇒
- if is_sep a0 then 〈parmove2,null_action ? n〉
- else match nth dst ? a (None ?) with
- [ None ⇒ 〈parmove2,null_action ? n〉
- | Some a1 ⇒ 〈parmove1,change_vec ? (S n)
- (change_vec ?(S n)
- (null_action ? n) (〈None sig,D〉) src)
- (〈None ?,D〉) dst〉 ] ]
- | S q ⇒ match q with
- [ O ⇒ (* 1 *) 〈parmove1,null_action ? n〉
- | S _ ⇒ (* 2 *) 〈parmove2,null_action ? n〉 ] ].
-
-definition parmove_step ≝
- λsrc,dst,sig,n,D,is_sep.
- mk_mTM sig n parmove_states (trans_parmove_step src dst sig n D is_sep)
- parmove0 (λq.q == parmove1 ∨ q == parmove2).
-
-definition R_parmove_step_true ≝
- λsrc,dst,sig,n,D,is_sep.λint,outt: Vector (tape sig) (S n).
- ∃x1,x2.
- current ? (nth src ? int (niltape ?)) = Some ? x1 ∧
- current ? (nth dst ? int (niltape ?)) = Some ? x2 ∧
- is_sep x1 = false ∧
- outt = change_vec ??
- (change_vec ?? int
- (tape_move_mono ? (nth src ? int (niltape ?)) (〈None ?,D〉)) src)
- (tape_move_mono ? (nth dst ? int (niltape ?)) (〈None ?,D〉)) dst.
-
-definition R_parmove_step_false ≝
- λsrc,dst:nat.λsig,n,is_sep.λint,outt: Vector (tape sig) (S n).
- ((∃x1.
- current ? (nth src ? int (niltape ?)) = Some ? x1 ∧
- is_sep x1 = true) ∨
- current ? (nth src ? int (niltape ?)) = None ? ∨
- current ? (nth dst ? int (niltape ?)) = None ?) ∧
- outt = int.
-
-lemma parmove_q0_q2_null_src :
- ∀src,dst,sig,n,D,is_sep,v.src < S n → dst < S n →
- nth src ? (current_chars ?? v) (None ?) = None ? →
- step sig n (parmove_step src dst sig n D is_sep)
- (mk_mconfig ??? parmove0 v) =
- mk_mconfig ??? parmove2 v.
-#src #dst #sig #n #D #is_sep #v #Hsrc #Hdst #Hcurrent
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?);
-@eq_f2
-[ whd in ⊢ (??(???%)?); >Hcurrent %
-| whd in ⊢ (??(????(???%))?); >Hcurrent @tape_move_null_action ]
-qed.
-
-lemma parmove_q0_q2_sep :
- ∀src,dst,sig,n,D,is_sep,v,s.src < S n → dst < S n →
- nth src ? (current_chars ?? v) (None ?) = Some ? s → is_sep s = true →
- step sig n (parmove_step src dst sig n D is_sep)
- (mk_mconfig ??? parmove0 v) =
- mk_mconfig ??? parmove2 v.
-#src #dst #sig #n #D #is_sep #v #s #Hsrc #Hdst #Hcurrent #Hsep
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?);
-@eq_f2
-[ whd in ⊢ (??(???%)?); >Hcurrent whd in ⊢ (??(???%)?); >Hsep %
-| whd in ⊢ (??(????(???%))?); >Hcurrent
- whd in ⊢ (??(????(???%))?); >Hsep @tape_move_null_action ]
-qed.
-
-lemma parmove_q0_q2_null_dst :
- ∀src,dst,sig,n,D,is_sep,v,s.src < S n → dst < S n →
- nth src ? (current_chars ?? v) (None ?) = Some ? s → is_sep s = false →
- nth dst ? (current_chars ?? v) (None ?) = None ? →
- step sig n (parmove_step src dst sig n D is_sep)
- (mk_mconfig ??? parmove0 v) =
- mk_mconfig ??? parmove2 v.
-#src #dst #sig #n #D #is_sep #v #s #Hsrc #Hdst #Hcursrc #Hsep #Hcurdst
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?);
-@eq_f2
-[ whd in ⊢ (??(???%)?); >Hcursrc whd in ⊢ (??(???%)?); >Hsep >Hcurdst %
-| whd in ⊢ (??(????(???%))?); >Hcursrc
- whd in ⊢ (??(????(???%))?); >Hsep >Hcurdst @tape_move_null_action ]
-qed.
-
-lemma parmove_q0_q1 :
- ∀src,dst,sig,n,D,is_sep,v.src ≠ dst → src < S n → dst < S n →
- ∀a1,a2.
- nth src ? (current_chars ?? v) (None ?) = Some ? a1 →
- nth dst ? (current_chars ?? v) (None ?) = Some ? a2 →
- is_sep a1 = false →
- step sig n (parmove_step src dst sig n D is_sep)
- (mk_mconfig ??? parmove0 v) =
- mk_mconfig ??? parmove1
- (change_vec ? (S n)
- (change_vec ?? v
- (tape_move_mono ? (nth src ? v (niltape ?)) (〈None ?, D〉)) src)
- (tape_move_mono ? (nth dst ? v (niltape ?)) (〈None ?, D〉)) dst).
-#src #dst #sig #n #D #is_sep #v #Hneq #Hsrc #Hdst
-#a1 #a2 #Hcursrc #Hcurdst #Hsep
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?); @eq_f2
-[ whd in match (trans ????);
- >Hcursrc >Hcurdst whd in ⊢ (??(???%)?); >Hsep //
-| whd in match (trans ????);
- >Hcursrc >Hcurdst whd in ⊢ (??(????(???%))?); >Hsep whd in ⊢ (??(????(???%))?);
- <(change_vec_same ?? v dst (niltape ?)) in ⊢ (??%?);
- >tape_move_multi_def >pmap_change
- <(change_vec_same ?? v src (niltape ?)) in ⊢ (??%?);
- >pmap_change <tape_move_multi_def >tape_move_null_action
- @eq_f2 // @eq_f2 // >nth_change_vec_neq //
-]
-qed.
-
-lemma sem_parmove_step :
- ∀src,dst,sig,n,D,is_sep.src ≠ dst → src < S n → dst < S n →
- parmove_step src dst sig n D is_sep ⊨
- [ parmove1: R_parmove_step_true src dst sig n D is_sep,
- R_parmove_step_false src dst sig n is_sep ].
-#src #dst #sig #n #D #is_sep #Hneq #Hsrc #Hdst #int
-lapply (refl ? (current ? (nth src ? int (niltape ?))))
-cases (current ? (nth src ? int (niltape ?))) in ⊢ (???%→?);
-[ #Hcursrc %{2} %
- [| % [ %
- [ whd in ⊢ (??%?); >parmove_q0_q2_null_src /2/
- | normalize in ⊢ (%→?); #H destruct (H) ]
- | #_ % // % %2 // ] ]
-| #a #Ha cases (true_or_false (is_sep a)) #Hsep
- [ %{2} %
- [| % [ %
- [ whd in ⊢ (??%?); >(parmove_q0_q2_sep … Hsep) /2/
- | normalize in ⊢ (%→?); #H destruct (H) ]
- | #_ % // % % %{a} % // ] ]
- | lapply (refl ? (current ? (nth dst ? int (niltape ?))))
- cases (current ? (nth dst ? int (niltape ?))) in ⊢ (???%→?);
- [ #Hcurdst %{2} %
- [| % [ %
- [ whd in ⊢ (??%?); >(parmove_q0_q2_null_dst … Hsep) /2/
- | normalize in ⊢ (%→?); #H destruct (H) ]
- | #_ % // %2 // ] ]
- | #b #Hb %{2} %
- [| % [ %
- [whd in ⊢ (??%?); >(parmove_q0_q1 … Hneq Hsrc Hdst ? b ?? Hsep) //
- | #_ %{a} %{b} % // % // % // ]
- | * #H @False_ind @H % ]
-]]]]
-qed.
-
-definition parmove ≝ λsrc,dst,sig,n,D,is_sep.
- whileTM … (parmove_step src dst sig n D is_sep) parmove1.
-
-definition R_parmoveL ≝
- λsrc,dst,sig,n,is_sep.λint,outt: Vector (tape sig) (S n).
- (∀ls,x,xs,rs,sep.
- nth src ? int (niltape ?) = midtape sig (xs@sep::ls) x rs →
- (∀c.memb ? c (x::xs) = true → is_sep c = false) → is_sep sep = true →
- ∀ls0,x0,target,c,rs0.|xs| = |target| →
- nth dst ? int (niltape ?) = midtape sig (target@c::ls0) x0 rs0 →
- outt = change_vec ??
- (change_vec ?? int (midtape sig ls sep (reverse ? xs@x::rs)) src)
- (midtape sig ls0 c (reverse ? target@x0::rs0)) dst) ∧
- (((∃s.current ? (nth src ? int (niltape ?)) = Some ? s ∧ is_sep s = true) ∨
- current ? (nth src ? int (niltape ?)) = None ? ∨
- current ? (nth dst ? int (niltape ?)) = None ?) →
- outt = int).
-
-lemma wsem_parmoveL : ∀src,dst,sig,n,is_sep.src ≠ dst → src < S n → dst < S n →
- parmove src dst sig n L is_sep ⊫ R_parmoveL src dst sig n is_sep.
-#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
-[ 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_src #ls #x0 #xs #rs #sep #Hsrctc >Hsrctc in Hcur_src;
- normalize in ⊢ (%→?); #H destruct (H)]
- |#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)
- ]
-| #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
- >Hsrc_tc in Hc0; normalize in ⊢ (%→?); #Hc0 destruct (Hc0)
- >Hdst_tc in Hd; >Hsrc_tc @(list_cases2 … Hlen)
- [ #Hxsnil #Htargetnil >Hxsnil >Htargetnil #Hd >IH2
- [2: %1 %1 %{sep} % // >Hd >nth_change_vec_neq [|@(sym_not_eq … Hneq)]
- >nth_change_vec //]
- >Hd -Hd @(eq_vec … (niltape ?))
- #i #Hi cases (decidable_eq_nat i src) #Hisrc
- [ >Hisrc >(nth_change_vec_neq … src dst) [|@(sym_not_eq … Hneq)]
- >nth_change_vec //
- | cases (decidable_eq_nat i dst) #Hidst
- [ >Hidst >nth_change_vec //
- | >nth_change_vec_neq [|@(sym_not_eq … Hidst)]
- >nth_change_vec_neq [|@(sym_not_eq … Hisrc)] %
- ]
- ]
- | #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 … ?? td ?? src dst) //
- >change_vec_change_vec
- >(change_vec_commute … ?? td ?? dst src) [|@sym_not_eq //]
- >change_vec_change_vec
- >reverse_cons >associative_append
- >reverse_cons >associative_append %
- | >Hd >nth_change_vec //
- | >Hxs in Hlen; >Htarget normalize #Hlen destruct (Hlen) //
- | <Hxs #c1 #Hc1 @Hnosep @memb_cons //
- | >Hd >nth_change_vec_neq [|@sym_not_eq //]
- >nth_change_vec // ]
- ]
- | >Hc0 >Hc1 * [* [ * #c * #Hc destruct (Hc) >Hc0nosep]] #Habs destruct (Habs)
- ] ]
-qed.
-
-lemma terminate_parmoveL : ∀src,dst,sig,n,is_sep,t.
- src ≠ dst → src < S n → dst < S n →
- parmove src dst sig n L is_sep ↓ t.
-#src #dst #sig #n #is_sep #t #Hneq #Hsrc #Hdst
-@(terminate_while … (sem_parmove_step …)) //
-<(change_vec_same … t src (niltape ?))
-cases (nth src (tape sig) t (niltape ?))
-[ % #t1 * #x1 * #x2 * * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct
-|2,3: #a0 #al0 % #t1 * #x1 * #x2 * * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct
-| #ls lapply t -t elim ls
- [#t #c #rs % #t1 * #x1 * #x2 * * * >nth_change_vec // normalize in ⊢ (%→?);
- #H1 destruct (H1) #Hcurdst #Hxsep >change_vec_change_vec #Ht1 %
- #t2 * #y1 * #y2 * * * >Ht1 >nth_change_vec_neq [|@sym_not_eq //]
- >nth_change_vec // normalize in ⊢ (%→?); #H destruct (H)
- |#l0 #ls0 #IH #t #c #rs % #t1 * #x1 * #x2 * * * >nth_change_vec //
- normalize in ⊢ (%→?); #H destruct (H) #Hcurdst #Hxsep
- >change_vec_change_vec >change_vec_commute // #Ht1 >Ht1 @IH
- ]
-]
-qed.
-
-lemma sem_parmoveL : ∀src,dst,sig,n,is_sep.
- src ≠ dst → src < S n → dst < S n →
- parmove src dst sig n L is_sep ⊨ R_parmoveL src dst sig n is_sep.
-#src #dst #sig #n #is_sep #Hneq #Hsrc #Hdst @WRealize_to_Realize
-[/2/ | @wsem_parmoveL //]
-qed.
\ No newline at end of file
+++ /dev/null
-(*
- ||M|| This file is part of HELM, an Hypertextual, Electronic
- ||A|| Library of Mathematics, developed at the Computer Science
- ||T|| Department of the University of Bologna, Italy.
- ||I||
- ||T||
- ||A||
- \ / This file is distributed under the terms of the
- \ / GNU General Public License Version 2
- V_____________________________________________________________*)
-
-include "turing/turing.ma".
-include "turing/inject.ma".
-include "turing/while_multi.ma".
-include "turing/while_machine.ma".
-include "turing/simple_machines.ma".
-include "turing/if_machine.ma".
-
-definition parmove_states ≝ initN 3.
-
-definition parmove0 : parmove_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
-definition parmove1 : parmove_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
-definition parmove2 : parmove_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
-
-(*
-
-src: a b c ... z ---→ a b c ... z
- ^ ^
-
-dst: _ _ _ ... _ ---→ a b c ... z
- ^ ^
-
-0) (x,_) → (x,_)(R,R) → 1
- (None,_) → None 2
-1) (_,_) → None 1
-2) (_,_) → None 2
-
-*)
-
-definition trans_parmove_step ≝
- λsrc,dst,sig,n,D.
- λp:parmove_states × (Vector (option sig) (S n)).
- let 〈q,a〉 ≝ p in
- match pi1 … q with
- [ O ⇒ match nth src ? a (None ?) with
- [ None ⇒ 〈parmove2,null_action sig n〉
- | Some a0 ⇒ match nth dst ? a (None ?) with
- [ None ⇒ 〈parmove2,null_action ? n〉
- | Some a1 ⇒ 〈parmove1,change_vec ? (S n)
- (change_vec ?(S n)
- (null_action ? n) (〈None ?,D〉) src)
- (〈None ?,D〉) dst〉 ] ]
- | S q ⇒ match q with
- [ O ⇒ (* 1 *) 〈parmove1,null_action ? n〉
- | S _ ⇒ (* 2 *) 〈parmove2,null_action ? n〉 ] ].
-
-definition parmove_step ≝
- λsrc,dst,sig,n,D.
- mk_mTM sig n parmove_states (trans_parmove_step src dst sig n D)
- parmove0 (λq.q == parmove1 ∨ q == parmove2).
-
-definition R_parmove_step_true ≝
- λsrc,dst,sig,n,D.λint,outt: Vector (tape sig) (S n).
- ∃x1,x2.
- current ? (nth src ? int (niltape ?)) = Some ? x1 ∧
- current ? (nth dst ? int (niltape ?)) = Some ? x2 ∧
- outt = change_vec ??
- (change_vec ?? int
- (tape_move ? (nth src ? int (niltape ?)) D) src)
- (tape_move ? (nth dst ? int (niltape ?)) D) dst.
-
-definition R_parmove_step_false ≝
- λsrc,dst:nat.λsig,n.λint,outt: Vector (tape sig) (S n).
- (current ? (nth src ? int (niltape ?)) = None ? ∨
- current ? (nth dst ? int (niltape ?)) = None ?) ∧
- outt = int.
-
-lemma parmove_q0_q2_null_src :
- ∀src,dst,sig,n,D,v.src < S n → dst < S n →
- nth src ? (current_chars ?? v) (None ?) = None ? →
- step sig n (parmove_step src dst sig n D)
- (mk_mconfig ??? parmove0 v) =
- mk_mconfig ??? parmove2 v.
-#src #dst #sig #n #D #v #Hsrc #Hdst #Hcurrent
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?);
-@eq_f2
-[ whd in ⊢ (??(???%)?); >Hcurrent %
-| whd in ⊢ (??(????(???%))?); >Hcurrent @tape_move_null_action ]
-qed.
-
-lemma parmove_q0_q2_null_dst :
- ∀src,dst,sig,n,D,v,s.src < S n → dst < S n →
- nth src ? (current_chars ?? v) (None ?) = Some ? s →
- nth dst ? (current_chars ?? v) (None ?) = None ? →
- step sig n (parmove_step src dst sig n D)
- (mk_mconfig ??? parmove0 v) =
- mk_mconfig ??? parmove2 v.
-#src #dst #sig #n #D #v #s #Hsrc #Hdst #Hcursrc #Hcurdst
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?);
-@eq_f2
-[ whd in ⊢ (??(???%)?); >Hcursrc whd in ⊢ (??(???%)?); >Hcurdst %
-| whd in ⊢ (??(????(???%))?); >Hcursrc
- whd in ⊢ (??(????(???%))?); >Hcurdst @tape_move_null_action ]
-qed.
-
-lemma parmove_q0_q1 :
- ∀src,dst,sig,n,D,v.src ≠ dst → src < S n → dst < S n →
- ∀a1,a2.
- nth src ? (current_chars ?? v) (None ?) = Some ? a1 →
- nth dst ? (current_chars ?? v) (None ?) = Some ? a2 →
- step sig n (parmove_step src dst sig n D)
- (mk_mconfig ??? parmove0 v) =
- mk_mconfig ??? parmove1
- (change_vec ? (S n)
- (change_vec ?? v
- (tape_move ? (nth src ? v (niltape ?)) D) src)
- (tape_move ? (nth dst ? v (niltape ?)) D) dst).
-#src #dst #sig #n #D #v #Hneq #Hsrc #Hdst
-#a1 #a2 #Hcursrc #Hcurdst
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?); @eq_f2
-[ whd in match (trans ????);
- >Hcursrc >Hcurdst %
-| whd in match (trans ????);
- >Hcursrc >Hcurdst whd in ⊢ (??(????(???%))?);
- >tape_move_multi_def <(change_vec_same ?? v dst (niltape ?)) in ⊢ (??%?);
- >pmap_change <(change_vec_same ?? v src (niltape ?)) in ⊢(??%?);
- >pmap_change <tape_move_multi_def >tape_move_null_action
- @eq_f2 // >nth_change_vec_neq //
-]
-qed.
-
-lemma sem_parmove_step :
- ∀src,dst,sig,n,D.src ≠ dst → src < S n → dst < S n →
- parmove_step src dst sig n D ⊨
- [ parmove1: R_parmove_step_true src dst sig n D,
- R_parmove_step_false src dst sig n ].
-#src #dst #sig #n #D #Hneq #Hsrc #Hdst #int
-lapply (refl ? (current ? (nth src ? int (niltape ?))))
-cases (current ? (nth src ? int (niltape ?))) in ⊢ (???%→?);
-[ #Hcursrc %{2} %
- [| % [ %
- [ whd in ⊢ (??%?); >parmove_q0_q2_null_src /2/
- | normalize in ⊢ (%→?); #H destruct (H) ]
- | #_ % // % // ] ]
-| #a #Ha lapply (refl ? (current ? (nth dst ? int (niltape ?))))
- cases (current ? (nth dst ? int (niltape ?))) in ⊢ (???%→?);
- [ #Hcurdst %{2} %
- [| % [ %
- [ whd in ⊢ (??%?); >(parmove_q0_q2_null_dst …) /2/
- | normalize in ⊢ (%→?); #H destruct (H) ]
- | #_ % // %2 // ] ]
- | #b #Hb %{2} %
- [| % [ %
- [whd in ⊢ (??%?); >(parmove_q0_q1 … Hneq Hsrc Hdst ? b ??)
- [2: <(nth_vec_map ?? (current …) dst ? int (niltape ?)) //
- |3: <(nth_vec_map ?? (current …) src ? int (niltape ?)) //
- | // ]
- | #_ %{a} %{b} % // % // ]
- | * #H @False_ind @H % ]
-]]]
-qed.
-
-definition parmove ≝ λsrc,dst,sig,n,D.
- whileTM … (parmove_step src dst sig n D) parmove1.
-
-definition R_parmoveL ≝
- λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
- (∀x,xs,rs.
- nth src ? int (niltape ?) = midtape sig xs x rs →
- ∀ls0,x0,target,rs0.|xs| = |target| →
- nth dst ? int (niltape ?) = midtape sig (target@ls0) x0 rs0 →
- outt = change_vec ??
- (change_vec ?? int (mk_tape sig [] (None ?) (reverse ? xs@x::rs)) src)
- (mk_tape sig (tail ? ls0) (option_hd ? ls0) (reverse ? target@x0::rs0)) dst) ∧
- (∀x,xs,rs.
- nth dst ? int (niltape ?) = midtape sig xs x rs →
- ∀ls0,x0,target,rs0.|xs| = |target| →
- nth src ? int (niltape ?) = midtape sig (target@ls0) x0 rs0 →
- outt = change_vec ??
- (change_vec ?? int (mk_tape sig [] (None ?) (reverse ? xs@x::rs)) dst)
- (mk_tape sig (tail ? ls0) (option_hd ? ls0) (reverse ? target@x0::rs0)) src) ∧
- ((current ? (nth src ? int (niltape ?)) = None ? ∨
- current ? (nth dst ? int (niltape ?)) = None ?) →
- outt = int).
-
-lemma wsem_parmoveL : ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n →
- parmove src dst sig n L ⊫ R_parmoveL src dst sig n.
-#src #dst #sig #n #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
-lapply (sem_while … (sem_parmove_step src dst sig n L Hneq Hsrc Hdst) … Hloop) //
--Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
-[ whd in ⊢ (%→?); * #H #Houtc % [2: #_ @Houtc ] cases H #Hcurtb
- [ %
- [ #x #xs #rs #Hsrctb >Hsrctb in Hcurtb; normalize in ⊢ (%→?);
- #Hfalse destruct (Hfalse)
- | #x #xs #rs #Hdsttb #ls0 #x0 #target #rs0 #Hlen #Hsrctb >Hsrctb in Hcurtb;
- normalize in ⊢ (%→?); #H destruct (H)
- ]
- | %
- [ #x #xs #rs #Hsrctb #ls0 #x0 #target
- #rs0 #Hlen #Hdsttb >Hdsttb in Hcurtb; normalize in ⊢ (%→?); #H destruct (H)
- | #x #xs #rs #Hdsttb >Hdsttb in Hcurtb; normalize in ⊢ (%→?);
- #Hfalse destruct (Hfalse)
- ]
- ]
-| #td #te * #c0 * #c1 * * #Hc0 #Hc1 #Hd #Hstar #IH #He
- lapply (IH He) -IH * * #IH1a #IH1b #IH2 % [ %
- [ #x #xs #rs #Hsrc_td #ls0 #x0 #target
- #rs0 #Hlen #Hdst_td
- >Hsrc_td in Hc0; normalize in ⊢ (%→?); #Hc0 destruct (Hc0)
- >Hdst_td in Hd; >Hsrc_td @(list_cases2 … Hlen)
- [ #Hxsnil #Htargetnil >Hxsnil >Htargetnil #Hd >IH2
- [2: %1 >Hd >nth_change_vec_neq [|@(sym_not_eq … Hneq)]
- >nth_change_vec //]
- >Hd -Hd @(eq_vec … (niltape ?))
- #i #Hi cases (decidable_eq_nat i src) #Hisrc
- [ >Hisrc >(nth_change_vec_neq … src dst) [|@(sym_not_eq … Hneq)]
- >nth_change_vec //
- >(nth_change_vec_neq … src dst) [|@(sym_not_eq … Hneq)]
- >nth_change_vec //
- | cases (decidable_eq_nat i dst) #Hidst
- [ >Hidst >nth_change_vec // >nth_change_vec //
- >Hdst_td in Hc1; >Htargetnil
- normalize in ⊢ (%→?); #Hc1 destruct (Hc1) cases ls0 //
- | >nth_change_vec_neq [|@(sym_not_eq … Hidst)]
- >nth_change_vec_neq [|@(sym_not_eq … Hisrc)]
- >nth_change_vec_neq [|@(sym_not_eq … Hidst)]
- >nth_change_vec_neq [|@(sym_not_eq … Hisrc)] %
- ]
- ]
- | #hd1 #hd2 #tl1 #tl2 #Hxs #Htarget >Hxs >Htarget #Hd
- >(IH1a hd1 tl1 (c0::rs) ? ls0 hd2 tl2 (x0::rs0))
- [ >Hd >(change_vec_commute … ?? td ?? src dst) //
- >change_vec_change_vec
- >(change_vec_commute … ?? td ?? dst src) [|@sym_not_eq //]
- >change_vec_change_vec
- >reverse_cons >associative_append
- >reverse_cons >associative_append %
- | >Hd >nth_change_vec //
- | >Hxs in Hlen; >Htarget normalize #Hlen destruct (Hlen) //
- | >Hd >nth_change_vec_neq [|@sym_not_eq //]
- >nth_change_vec // ]
- ]
- | #x #xs #rs #Hdst_td #ls0 #x0 #target
- #rs0 #Hlen #Hsrc_td
- >Hdst_td in Hc0; normalize in ⊢ (%→?); #Hc0 destruct (Hc0)
- >Hsrc_td in Hd; >Hdst_td @(list_cases2 … Hlen)
- [ #Hxsnil #Htargetnil >Hxsnil >Htargetnil #Hd >IH2
- [2: %2 >Hd >nth_change_vec //]
- >Hd -Hd @(eq_vec … (niltape ?))
- #i #Hi cases (decidable_eq_nat i dst) #Hidst
- [ >Hidst >(nth_change_vec_neq … dst src) //
- >nth_change_vec // >nth_change_vec //
- | cases (decidable_eq_nat i src) #Hisrc
- [ >Hisrc >nth_change_vec // >(nth_change_vec_neq …) [|@sym_not_eq //]
- >Hsrc_td in Hc1; >Htargetnil
- normalize in ⊢ (%→?); #Hc1 destruct (Hc1) >nth_change_vec //
- cases ls0 //
- | >nth_change_vec_neq [|@(sym_not_eq … Hidst)]
- >nth_change_vec_neq [|@(sym_not_eq … Hisrc)]
- >nth_change_vec_neq [|@(sym_not_eq … Hisrc)]
- >nth_change_vec_neq [|@(sym_not_eq … Hidst)] %
- ]
- ]
- | #hd1 #hd2 #tl1 #tl2 #Hxs #Htarget >Hxs >Htarget #Hd
- >(IH1b hd1 tl1 (x::rs) ? ls0 hd2 tl2 (x0::rs0))
- [ >Hd >(change_vec_commute … ?? td ?? dst src) [|@sym_not_eq //]
- >change_vec_change_vec
- >(change_vec_commute … ?? td ?? src dst) //
- >change_vec_change_vec
- >reverse_cons >associative_append
- >reverse_cons >associative_append
- >change_vec_commute [|@sym_not_eq //] %
- | >Hd >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec //
- | >Hxs in Hlen; >Htarget normalize #Hlen destruct (Hlen) //
- | >Hd >nth_change_vec // ]
- ]
- ]
-| >Hc0 >Hc1 * [ #Hc0 destruct (Hc0) | #Hc1 destruct (Hc1) ]
-] ]
-qed.
-
-lemma terminate_parmoveL : ∀src,dst,sig,n,t.
- src ≠ dst → src < S n → dst < S n →
- parmove src dst sig n L ↓ t.
-#src #dst #sig #n #t #Hneq #Hsrc #Hdst
-@(terminate_while … (sem_parmove_step …)) //
-<(change_vec_same … t src (niltape ?))
-cases (nth src (tape sig) t (niltape ?))
-[ % #t1 * #x1 * #x2 * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct
-|2,3: #a0 #al0 % #t1 * #x1 * #x2 * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct
-| #ls lapply t -t elim ls
- [#t #c #rs % #t1 * #x1 * #x2 * * >nth_change_vec // normalize in ⊢ (%→?);
- #H1 destruct (H1) #Hcurdst >change_vec_change_vec #Ht1 %
- #t2 * #y1 * #y2 * * >Ht1 >nth_change_vec_neq [|@sym_not_eq //]
- >nth_change_vec // normalize in ⊢ (%→?); #H destruct (H)
- |#l0 #ls0 #IH #t #c #rs % #t1 * #x1 * #x2 * * >nth_change_vec //
- normalize in ⊢ (%→?); #H destruct (H) #Hcurdst
- >change_vec_change_vec >change_vec_commute // #Ht1 >Ht1 @IH
- ]
-]
-qed.
-
-lemma sem_parmoveL : ∀src,dst,sig,n.
- src ≠ dst → src < S n → dst < S n →
- parmove src dst sig n L ⊨ R_parmoveL src dst sig n.
-#src #dst #sig #n #Hneq #Hsrc #Hdst @WRealize_to_Realize
-[/2/ | @wsem_parmoveL //]
-qed.
-
-(* while {
- if current != null
- then move_r
- else nop
- }
- *)
-
-definition mte_step ≝ λalpha,D.
-ifTM ? (test_null alpha) (single_finalTM ? (move alpha D)) (nop ?) tc_true.
-
-definition R_mte_step_true ≝ λalpha,D,t1,t2.
- ∃ls,c,rs.
- t1 = midtape alpha ls c rs ∧ t2 = tape_move ? t1 D.
-
-definition R_mte_step_false ≝ λalpha.λt1,t2:tape alpha.
- current ? t1 = None ? ∧ t1 = t2.
-
-definition mte_acc : ∀alpha,D.states ? (mte_step alpha D) ≝
-λalpha,D.(inr … (inl … (inr … start_nop))).
-
-lemma sem_mte_step :
- ∀alpha,D.mte_step alpha D ⊨
- [ mte_acc … : R_mte_step_true alpha D, R_mte_step_false alpha ] .
-#alpha #D #ta
-@(acc_sem_if_app ??????????? (sem_test_null …)
- (sem_move_single …) (sem_nop alpha) ??)
-[ #tb #tc #td * #Hcurtb
- lapply (refl ? (current ? tb)) cases (current ? tb) in ⊢ (???%→?);
- [ #H @False_ind >H in Hcurtb; * /2/ ]
- -Hcurtb #c #Hcurtb #Htb whd in ⊢ (%→?); #Htc whd
- cases (current_to_midtape … Hcurtb) #ls * #rs #Hmidtb
- %{ls} %{c} %{rs} % //
-| #tb #tc #td * #Hcurtb #Htb whd in ⊢ (%→?); #Htc whd % // ]
-qed.
-
-definition move_to_end ≝ λsig,D.whileTM sig (mte_step sig D) (mte_acc …).
-
-definition R_move_to_end_r ≝
- λsig,int,outt.
- (current ? int = None ? → outt = int) ∧
- ∀ls,c,rs.int = midtape sig ls c rs → outt = mk_tape ? (reverse ? rs@c::ls) (None ?) [ ].
-
-lemma wsem_move_to_end_r : ∀sig. move_to_end sig R ⊫ R_move_to_end_r sig.
-#sig #ta #k #outc #Hloop
-lapply (sem_while … (sem_mte_step sig R) … Hloop) //
--Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
-[ * #Hcurtb #Houtc % /2/ #ls #c #rs #Htb >Htb in Hcurtb; normalize in ⊢ (%→?); #H destruct (H)
-| #tc #td * #ls * #c * #rs * #Htc >Htc cases rs
- [ normalize in ⊢ (%→?); #Htd >Htd #Hstar #IH whd in ⊢ (%→?); #Hfalse
- lapply (IH Hfalse) -IH * #Htd1 #_ %
- [ normalize in ⊢ (%→?); #H destruct (H)
- | #ls0 #c0 #rs0 #H destruct (H) >Htd1 // ]
- | #r0 #rs0 whd in ⊢ (???%→?); #Htd >Htd #Hstar #IH whd in ⊢ (%→?); #Hfalse
- lapply (IH Hfalse) -IH * #_ #IH %
- [ normalize in ⊢ (%→?); #H destruct (H)
- | #ls1 #c1 #rs1 #H destruct (H) >reverse_cons >associative_append @IH % ] ] ]
-qed.
-
-lemma terminate_move_to_end_r : ∀sig,t.move_to_end sig R ↓ t.
-#sig #t @(terminate_while … (sem_mte_step sig R …)) //
-cases t
-[ % #t1 * #ls * #c * #rs * #H destruct
-|2,3: #a0 #al0 % #t1 * #ls * #c * #rs * #H destruct
-| #ls #c #rs lapply c -c lapply ls -ls elim rs
- [ #ls #c % #t1 * #ls0 * #c0 * #rs0 * #Hmid #Ht1 destruct %
- #t2 * #ls1 * #c1 * #rs1 * normalize in ⊢ (%→?); #H destruct
- | #r0 #rs0 #IH #ls #c % #t1 * #ls1 * #c1 * #rs1 * #Hmid #Ht1 destruct @IH
- ]
-]
-qed.
-
-lemma sem_move_to_end_r : ∀sig. move_to_end sig R ⊨ R_move_to_end_r sig.
-#sig @WRealize_to_Realize //
-qed.
-
-definition R_move_to_end_l ≝
- λsig,int,outt.
- (current ? int = None ? → outt = int) ∧
- ∀ls,c,rs.int = midtape sig ls c rs → outt = mk_tape ? [ ] (None ?) (reverse ? ls@c::rs).
-
-lemma wsem_move_to_end_l : ∀sig. move_to_end sig L ⊫ R_move_to_end_l sig.
-#sig #ta #k #outc #Hloop
-lapply (sem_while … (sem_mte_step sig L) … Hloop) //
--Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
-[ * #Hcurtb #Houtc % /2/ #ls #c #rs #Htb >Htb in Hcurtb; normalize in ⊢ (%→?); #H destruct (H)
-| #tc #td * #ls * #c * #rs * #Htc >Htc cases ls
- [ normalize in ⊢ (%→?); #Htd >Htd #Hstar #IH whd in ⊢ (%→?); #Hfalse
- lapply (IH Hfalse) -IH * #Htd1 #_ %
- [ normalize in ⊢ (%→?); #H destruct (H)
- | #ls0 #c0 #rs0 #H destruct (H) >Htd1 // ]
- | #l0 #ls0 whd in ⊢ (???%→?); #Htd >Htd #Hstar #IH whd in ⊢ (%→?); #Hfalse
- lapply (IH Hfalse) -IH * #_ #IH %
- [ normalize in ⊢ (%→?); #H destruct (H)
- | #ls1 #c1 #rs1 #H destruct (H) >reverse_cons >associative_append @IH % ] ] ]
-qed.
-
-lemma terminate_move_to_end_l : ∀sig,t.move_to_end sig L ↓ t.
-#sig #t @(terminate_while … (sem_mte_step sig L …)) //
-cases t
-[ % #t1 * #ls * #c * #rs * #H destruct
-|2,3: #a0 #al0 % #t1 * #ls * #c * #rs * #H destruct
-| #ls elim ls
- [ #c #rs % #t1 * #ls0 * #c0 * #rs0 * #Hmid #Ht1 destruct %
- #t2 * #ls1 * #c1 * #rs1 * normalize in ⊢ (%→?); #H destruct
- | #l0 #ls0 #IH #c #rs % #t1 * #ls1 * #c1 * #rs1 * #Hmid #Ht1 destruct @IH
- ]
-]
-qed.
-
-lemma sem_move_to_end_l : ∀sig. move_to_end sig L ⊨ R_move_to_end_l sig.
-#sig @WRealize_to_Realize //
-qed.
+++ /dev/null
-(*
- ||M|| This file is part of HELM, an Hypertextual, Electronic
- ||A|| Library of Mathematics, developed at the Computer Science
- ||T|| Department of the University of Bologna, Italy.
- ||I||
- ||T||
- ||A||
- \ / This file is distributed under the terms of the
- \ / GNU General Public License Version 2
- V_____________________________________________________________*)
-
-include "turing/turing.ma".
-include "turing/inject.ma".
-include "turing/while_multi.ma".
-
-definition partest_states ≝ initN 3.
-
-definition partest0 : partest_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
-definition partest1 : partest_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
-definition partest2 : partest_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
-
-definition trans_partest ≝
- λsig,n,test.
- λp:partest_states × (Vector (option sig) (S n)).
- let 〈q,a〉 ≝ p in
- if test a then 〈partest1,null_action sig n〉
- else 〈partest2,null_action ? n〉.
-
-definition partest ≝
- λsig,n,test.
- mk_mTM sig n partest_states (trans_partest sig n test)
- partest0 (λq.q == partest1 ∨ q == partest2).
-
-definition R_partest_true ≝
- λsig,n,test.λint,outt: Vector (tape sig) (S n).
- test (current_chars ?? int) = true ∧ outt = int.
-
-definition R_partest_false ≝
- λsig,n,test.λint,outt: Vector (tape sig) (S n).
- test (current_chars ?? int) = false ∧ outt = int.
-
-lemma partest_q0_q1:
- ∀sig,n,test,v.
- test (current_chars ?? v) = true →
- step sig n (partest sig n test)(mk_mconfig ??? partest0 v)
- = mk_mconfig ??? partest1 v.
-#sig #n #test #v #Htest
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?);
-@eq_f2
-[ whd in ⊢ (??(???%)?); >Htest %
-| whd in ⊢ (??(????(???%))?); >Htest @tape_move_null_action ]
-qed.
-
-lemma partest_q0_q2:
- ∀sig,n,test,v.
- test (current_chars ?? v) = false →
- step sig n (partest sig n test)(mk_mconfig ??? partest0 v)
- = mk_mconfig ??? partest2 v.
-#sig #n #test #v #Htest
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?);
-@eq_f2
-[ whd in ⊢ (??(???%)?); >Htest %
-| whd in ⊢ (??(????(???%))?); >Htest @tape_move_null_action ]
-qed.
-
-lemma sem_partest:
- ∀sig,n,test.
- partest sig n test ⊨
- [ partest1: R_partest_true sig n test, R_partest_false sig n test ].
-#sig #n #test #int
-cases (true_or_false (test (current_chars ?? int))) #Htest
-[ %{2} %{(mk_mconfig ? partest_states n partest1 int)} %
- [ % [ whd in ⊢ (??%?); >partest_q0_q1 /2/ | #_ % // ]
- | * #H @False_ind @H %
- ]
-| %{2} %{(mk_mconfig ? partest_states n partest2 int)} %
- [ % [ whd in ⊢ (??%?); >partest_q0_q2 /2/
- | whd in ⊢ (??%%→?); #H destruct (H)]
- | #_ % //]
-]
-qed.
\ No newline at end of file
#Heqout <Heqout in tuplet2; <Heqq <Heqa >tuplet1
@append_l2_injective %
qed.
-
-lemma cfg_in_table_to_tuple: ∀n,l,h,c. is_config n c →
- ∀ll,lr.table_TM n l h = ll@c@lr →
- ∃out,m,lr0. lr = out@m::lr0 ∧ is_config n (bar::out) ∧ m ≠ bar.
-#n #l #h #c * #qin * #cin * * * #H1 #H2 #H3 #H4
-#ll #lr lapply ll -ll elim l
- [>H4 #ll cases ll normalize [|#hd #tl ] #Habs destruct
- |#t1 #othert #Hind #ll >table_TM_cons #Htuple
- cut (S n < |ll@c@lr|)
- [<Htuple >length_append >(length_of_tuple … (is_tuple … ))
- /2 by transitive_lt, le_n/] #Hsplit lapply Htuple -Htuple
- cases (is_tuple … n h t1) #q1 * #c1 * #q2 * #c2 * #m
- * * * * * * * #Hq1 #Hq2 #Hc1 #Hc2 #Hm #Hlen1 #Hlen2
- whd in ⊢ (???%→?); #Ht1
- (* if ll is empty we match the first tuple t1, otherwise
- we match inside othert *)
- cases ll
- [>H4 >Ht1 normalize in ⊢ (???%→?);
- >associative_append whd in ⊢ (??%?→?); #Heq destruct (Heq) -Heq
- >associative_append in e0; #e0
- lapply (append_l1_injective … e0) [>H3 @Hlen1] #Heq1
- lapply (append_l2_injective … e0) [>H3 @Hlen1]
- normalize in ⊢ (???%→?); whd in ⊢ (??%?→?); #Htemp
- lapply (cons_injective_l ????? Htemp) #Hc1
- lapply (cons_injective_r ????? Htemp) -Htemp #Heq2
- %{(q2@[c2])} %{m} %{(table_TM n othert h)} % // %
- [ <Heq2 >associative_append >associative_append % | %{q2} %{c2} % // % // % // ]
- |(* ll not nil *)
- #b #tl >Ht1 normalize in ⊢ (???%→?);
- whd in ⊢ (??%?→?); #Heq destruct (Heq)
- cases (compare_append … e0) #l *
- [* cases l
- [#_ #Htab cases (Hind [ ] (sym_eq … Htab)) #out * #m0 * #lr0 * * #Hlr #Hcfg #Hm0
- %{out} %{m0} %{lr0} % // % //
- |(* this case is absurd *)
- #al #tll #Heq1 >H4 #Heq2 @False_ind
- lapply (cons_injective_l ? bar … Heq2) #Hbar <Hbar in Heq1; #Heq1
- @(absurd (mem ? bar (q1@(c1::q2@[c2; m]))))
- [>Heq1 @mem_append_l2 %1 //
- |% #Hmembar cases (mem_append ???? Hmembar) -Hmembar
- [#Hmembar lapply(Hq1 bar Hmembar) normalize #Habs destruct (Habs)
- |* [#Habs @absurd //]
- #Hmembar cases (mem_append ???? Hmembar) -Hmembar
- [#Hmembar lapply(Hq2 bar Hmembar) normalize #Habs destruct (Habs)
- |* [#Habs @absurd //] #Hmembar @(absurd ?? Hm) @sym_eq @mem_single //
- ]
- ]
- ]
- ]
- |* #Htl #Htab cases (Hind … Htab) #out * #m0 * #lr0 * * #Hlr #Hcfg #Hm0
- %{out} %{m0} %{lr0} % // % //
- ]
- ]
- ]
-qed.
-
include "turing/multi_universal/unistep_aux.ma".
+include "turing/multi_universal/match.ma".
definition exec_move ≝
cfg_to_obj · tape_move_obj · restart_tape prg 2 · obj_to_cfg.
\ / GNU General Public License Version 2
V_____________________________________________________________*)
-include "turing/multi_universal/moves_2.ma".
-include "turing/multi_universal/match.ma".
-include "turing/multi_universal/copy.ma".
+include "turing/auxiliary_machines.ma".
+include "turing/auxiliary_multi_machines.ma".
include "turing/multi_universal/alphabet.ma".
include "turing/multi_universal/tuples.ma".
\ / GNU General Public License Version 2
V_____________________________________________________________*)
-include "turing/simple_machines.ma".
+include "turing/auxiliary_multi_machines.ma".
include "turing/multi_universal/unistep.ma".
definition stop ≝ λc.