+++ /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_____________________________________________________________*)
-
-
-(* ALphabet of the universal machine *)
-
-include "basics/finset.ma".
-
-(*
-include "turing/if_machine.ma".
-include "turing/universal/tests.ma". *)
-
-inductive unialpha : Type[0] ≝
-| bit : bool → unialpha
-| null : unialpha
-| comma : unialpha
-| bar : unialpha
-| grid : unialpha.
-
-definition unialpha_eq ≝
- λa1,a2.match a1 with
- [ bit x ⇒ match a2 with [ bit y ⇒ ¬ xorb x y | _ ⇒ false ]
- | null ⇒ match a2 with [ null ⇒ true | _ ⇒ false ]
- | comma ⇒ match a2 with [ comma ⇒ true | _ ⇒ false ]
- | bar ⇒ match a2 with [ bar ⇒ true | _ ⇒ false ]
- | grid ⇒ match a2 with [ grid ⇒ true | _ ⇒ false ] ].
-
-definition DeqUnialpha ≝ mk_DeqSet unialpha unialpha_eq ?.
-* [ #x * [ #y cases x cases y normalize % // #Hfalse destruct
- | *: normalize % #Hfalse destruct ]
- |*: * [1,6,11,16: #y ] normalize % #H1 destruct % ]
-qed.
-
-lemma unialpha_unique :
- uniqueb DeqUnialpha [bit true;bit false;null;comma;bar;grid] = true.
-// qed.
-
-lemma unialpha_complete :∀x:DeqUnialpha.
- memb ? x [bit true;bit false;null;comma;bar;grid] = true.
-* // * //
-qed.
-
-definition FSUnialpha ≝
- mk_FinSet DeqUnialpha [bit true;bit false;null;comma;bar;grid]
- unialpha_unique unialpha_complete.
-
-(*************************** testing characters *******************************)
-definition is_bit ≝ λc.match c with [ bit _ ⇒ true | _ ⇒ false ].
-
-definition is_null ≝ λc.match c with [ null ⇒ true | _ ⇒ false ].
-
-definition is_grid ≝ λc.match c with [ grid ⇒ true | _ ⇒ false ].
-
-definition is_bar ≝ λc.match c with [ bar ⇒ true | _ ⇒ false ].
-
-definition is_comma ≝ λc.match c with [ comma ⇒ true | _ ⇒ false ].
-
-definition bit_or_null ≝ λc.is_bit c ∨ is_null c.
-
-lemma bit_not_grid: ∀d. is_bit d = true → is_grid d = false.
-* // normalize #H destruct
-qed.
-
-lemma bit_or_null_not_grid: ∀d. bit_or_null d = true → is_grid d = false.
-* // normalize #H destruct
-qed.
-
-lemma bit_not_bar: ∀d. is_bit d = true → is_bar d = false.
-* // normalize #H destruct
-qed.
-
-lemma bit_or_null_not_bar: ∀d. bit_or_null d = true → is_bar d = false.
-* // normalize #H destruct
-qed.
-
-lemma is_bit_to_bit_or_null :
- ∀x.is_bit x = true → bit_or_null x = true.
-* // normalize #H destruct
-qed.
-
-(**************************** testing strings *********************************)
-definition is_marked ≝ λalpha.λp:FinProd … alpha FinBool.
- let 〈x,b〉 ≝ p in b.
-
-definition STape ≝ FinProd … FSUnialpha FinBool.
-
-definition only_bits ≝ λl.
- ∀c.memb STape c l = true → is_bit (\fst c) = true.
-
-definition only_bits_or_nulls ≝ λl.
- ∀c.memb STape c l = true → bit_or_null (\fst c) = true.
-
-definition no_grids ≝ λl.
- ∀c.memb STape c l = true → is_grid (\fst c) = false.
-
-definition no_bars ≝ λl.
- ∀c.memb STape c l = true → is_bar (\fst c) = false.
-
-definition no_marks ≝ λl.
- ∀c.memb STape c l = true → is_marked ? c = false.
-
-definition bar_or_grid ≝ λc:STape.is_bar (\fst c) ∨ is_grid (\fst c).
-
+++ /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/universal/marks.ma".
-
-definition copy_step_subcase ≝
- λalpha,c,elseM.ifTM ? (test_char ? (λx.x == 〈c,true〉))
- (seq (FinProd alpha FinBool) (adv_mark_r …)
- (seq ? (move_l …)
- (seq ? (adv_to_mark_l … (is_marked alpha))
- (seq ? (write ? 〈c,false〉)
- (seq ? (move_r …)
- (seq ? (mark …)
- (seq ? (move_r …) (adv_to_mark_r … (is_marked alpha)))))))))
- elseM tc_true.
-
-definition R_copy_step_subcase ≝
- λalpha,c,RelseM,t1,t2.
- ∀a,l1,x0,a0,l2,x,l3.
- t1 = midtape (FinProd … alpha FinBool) (l1@〈a0,false〉::〈x0,true〉::l2)
- 〈x,true〉 (〈a,false〉::l3) →
- (∀c.memb ? c l1 = true → is_marked ? c = false) →
- (x = c ∧ t2 = midtape ? (〈x,false〉::l1@〈a0,true〉::〈x,false〉::l2) 〈a,true〉 l3) ∨
- (x ≠ c ∧ RelseM t1 t2).
-
-lemma sem_copy_step_subcase :
- ∀alpha,c,elseM,RelseM. Realize ? elseM RelseM →
- Realize ? (copy_step_subcase alpha c elseM) (R_copy_step_subcase alpha c RelseM).
-#alpha #c #elseM #RelseM #HelseM #intape
-cases (sem_if ? (test_char ? (λx. x == 〈c,true〉)) ?????? tc_true (sem_test_char ? (λx.x == 〈c,true〉))
- (sem_seq ????? (sem_adv_mark_r alpha)
- (sem_seq ????? (sem_move_l …)
- (sem_seq ????? (sem_adv_to_mark_l … (is_marked alpha))
- (sem_seq ????? (sem_write ? 〈c,false〉)
- (sem_seq ????? (sem_move_r …)
- (sem_seq ????? (sem_mark …)
- (sem_seq ????? (sem_move_r …) (sem_adv_to_mark_r … (is_marked alpha)))))))))
- HelseM intape)
-#k * #outc * #Hloop #HR %{k} %{outc} % [@Hloop] -Hloop
-#a #l1 #x0 #a0 #l2 #x #l3 #Hintape #Hl1marks cases HR -HR
-[ * #ta * whd in ⊢ (%→?); >Hintape * * #c0 * whd in ⊢ (??%?→?); #Hx #Hc #Hta
- * #tb * whd in ⊢ (%→?); * #Htb cases (Htb (l1@〈a0,false〉::〈x0,true〉::l2) x) -Htb
- #Htb lapply (Htb … Hta) -Htb #Htb #_ #_
- * #tc * whd in ⊢ (%→?); * #_ #Htc lapply (Htc … Htb) -Htb -Htc #Htc
- * #td * whd in ⊢ (%→?); * #_ #Htd cases (Htd … Htc) -Htd #_ #Htd cases (Htd (refl ??))
- -Htd #Htd lapply (Htd (l1@[〈a0,false〉]) 〈x0,true〉 l2 ???) //
- [#x1 #Hx1 cases (memb_append … Hx1) [ @Hl1marks | #Hsingle >(memb_single … Hsingle) % ]
- |whd in ⊢ (??%?); // ]
- -Htd #Htd #_
- * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd) -Hte -Htd -Htc #Hte
- * #tf * whd in ⊢ (%→?); * #_ #Htf lapply (Htf … Hte) -Hte -Htf >reverse_append #Htf
- * #tg * whd in ⊢ (%→?); * #Htg #_ lapply (Htg … Htf) -Htf -Htg >reverse_single #Htg
- * #th * whd in ⊢ (%→?); * #_ #Hth lapply (Hth … Htg) -Htg -Hth
- generalize in match Hl1marks; -Hl1marks @(list_elim_left … l1)
- [ #Hl1marks #Hth whd in ⊢ (%→?); * #_ #Houtc cases (Houtc … Hth) -Houtc
- [ * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
- * * #_ #Houtc #_ lapply (Houtc [] ?? (refl ??) (refl ??) Hl1marks) -Houtc
- #Houtc % >(\P Hc) in Hx; #Hx destruct (Hx) % // @Houtc
- | -l1 #c1 #l1 #_ #Hl1marks >reverse_append >reverse_single
- #Hth whd in ⊢ (%→?); * #_ #Houtc cases (Houtc … Hth) -Houtc
- [ * >Hl1marks [ #Hfalse destruct (Hfalse) ] @memb_append_l2 @memb_hd ]
- * * #_ #Houtc lapply (Houtc (reverse ? l1@[〈x,false〉]) 〈a,true〉 l3 ? (refl ??) ?) -Houtc
- [ #x1 #Hx1 cases (memb_append … Hx1) -Hx1 #Hx1
- [@Hl1marks @memb_append_l1 <(reverse_reverse … l1) @memb_reverse @Hx1
- |>(memb_single … Hx1) % ]
- | normalize >associative_append % ]
- #Houtc #_ % destruct (Hx) lapply (\P Hc) -Hc #Hc destruct (Hc) % //
- >Houtc >reverse_append >reverse_reverse >associative_append >associative_append % ]
-| * #ta * whd in ⊢ (%→?); >Hintape * #Hxc #Hta #Helse %2 %
- [| <Hta @Helse ]
- % #Hxc0 >Hxc0 in Hxc; #Hxc lapply (Hxc 〈c,true〉 (refl …)) #Hfalse
- cases (\Pf Hfalse) #Hfalse0 @Hfalse0 %
-]
-qed.
-
-(*
-if current = 0,tt
- then advance_mark_r;
- move_l;
- advance_to_mark_l;
- write(0,ff)
- move_r;
- mark;
- move_r;
- advance_to_mark_r;
-else if current = 1,tt
- then advance_mark_r;
- move_l;
- advance_to_mark_l;
- write(1,ff)
- move_r;
- mark;
- move_r;
- advance_to_mark_r;
-else if current = null
- then advance_mark_r;
- move_l;
- advance_to_mark_l
- adv_mark_r;
- move_r;
- advance_to_mark_r
-*)
-
-definition nocopy_subcase ≝
- ifTM STape (test_char ? (λx:STape.x == 〈null,true〉))
- (seq ? (adv_mark_r …)
- (seq ? (move_l …)
- (seq ? (adv_to_mark_l … (is_marked ?))
- (seq ? (adv_mark_r …)
- (seq ? (move_r …) (adv_to_mark_r … (is_marked ?)))))))
- (nop ?) tc_true.
-
-definition R_nocopy_subcase ≝
- λt1,t2.
- ∀a,l1,x0,a0,l2,x,l3.
- t1 = midtape STape (l1@〈a0,false〉::〈x0,true〉::l2)
- 〈x,true〉 (〈a,false〉::l3) →
- (∀c.memb ? c l1 = true → is_marked ? c = false) →
- (x = null ∧
- t2 = midtape ? (〈x,false〉::l1@〈a0,true〉::〈x0,false〉::l2) 〈a,true〉 l3) ∨
- (x ≠ null ∧ t2 = t1).
-
-lemma sem_nocopy_subcase : Realize ? nocopy_subcase R_nocopy_subcase.
-#intape
-cases (sem_if ? (test_char ? (λx:STape.x == 〈null,true〉)) ?????? tc_true
- (sem_test_char ? (λx:STape.x == 〈null,true〉))
- (sem_seq … (sem_adv_mark_r …)
- (sem_seq … (sem_move_l …)
- (sem_seq … (sem_adv_to_mark_l … (is_marked ?))
- (sem_seq … (sem_adv_mark_r …)
- (sem_seq … (sem_move_r …)
- (sem_adv_to_mark_r … (is_marked ?))))))) (sem_nop ?) intape)
-#k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#a #l1 #x0 #a0 #l2 #x #l3 #Hintape #Hl1marks cases HR -HR
-[ * #ta * whd in ⊢ (%→?); >Hintape * * #c * whd in ⊢ (??%?→?); #Hc destruct (Hc) #Hx #Hta
- * #tb * whd in ⊢ (%→?); * #Htb #_ cases (Htb (l1@〈a0,false〉::〈x0,true〉::l2) x) -Htb #Htb #_ lapply (Htb … Hta) -Hta -Htb #Htb
- * #tc * whd in ⊢ (%→?); * #_ #Htc lapply (Htc … Htb) -Htb -Htc #Htc
- * #td * whd in ⊢ (%→?); * #_ #Htd cases (Htd … Htc) -Htd #_ #Htd cases (Htd (refl …)) -Htd #Htd #_
- lapply (Htd (l1@[〈a0,false〉]) 〈x0,true〉 l2 ? (refl …) ?)
- [#x1 #Hx1 cases (memb_append … Hx1) [@Hl1marks| -Hx1 #Hx1 >(memb_single … Hx1) % ]
- |>associative_append % ] -Htd >reverse_append in ⊢ (???%→?); >associative_append in ⊢ (???%→?); #Htd
- * #te * whd in ⊢ (%→?); * #Hte cases (Hte l2 x0) -Hte #Hte #_ #_ lapply (Hte … Htd) -Hte -Htd -Htc #Hte
- * #tf * whd in ⊢ (%→?); * #_ #Htf lapply (Htf … Hte) -Hte -Htf
- generalize in match Hl1marks; -Hl1marks @(list_elim_left … l1)
- [ #Hl1marks #Hth whd in ⊢ (%→?); * #_ #Houtc cases (Houtc … Hth) -Houtc
- [ * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
- * * #_ #Houtc lapply (Houtc [] ?? (refl ??) (refl ??) Hl1marks) -Houtc
- #Houtc lapply (\P Hx) -Hx #Hx destruct (Hx) #_ % % [%] @Houtc
- | -l1 #c1 #l1 #_ #Hl1marks >reverse_append >reverse_single
- #Hth whd in ⊢ (%→?); * #_ #Houtc cases (Houtc … Hth) -Houtc
- [ * >Hl1marks [ #Hfalse destruct (Hfalse) ] @memb_append_l2 @memb_hd ]
- * * #Hc1 #Houtc #_ lapply (Houtc (reverse ? l1@[〈x,false〉]) 〈a,true〉 l3 ? (refl ??) ?) -Houtc
- [ #x1 #Hx1 cases (memb_append … Hx1) -Hx1 #Hx1
- [@Hl1marks @memb_append_l1 <(reverse_reverse … l1) @memb_reverse @Hx1
- |>(memb_single … Hx1) % ]
- | normalize >associative_append % ]
- #Houtc lapply (\P Hx) -Hx #Hx destruct (Hx) % % [%] >Houtc
- >reverse_append >reverse_reverse >associative_append >associative_append % ]
-| * #ta * whd in ⊢ (%→?); >Hintape * #Hxc
- cut (x ≠ null) [ % #Hx cases (\Pf (Hxc ? (refl …))) #Hfalse @Hfalse >Hx % ] -Hxc #Hxnull
- #Hta whd in ⊢ (%→?); #Houtc %2 % // <Hta @Houtc ]
-qed.
-
-definition copy_step ≝
- ifTM ? (test_char STape (λc.bit_or_null (\fst c)))
- (single_finalTM ? (copy_step_subcase FSUnialpha (bit false)
- (copy_step_subcase FSUnialpha (bit true) nocopy_subcase)))
- (nop ?)
- tc_true.
-
-definition R_copy_step_true ≝
- λt1,t2.
- ∀ls,c,rs. t1 = midtape STape ls 〈c,true〉 rs →
- bit_or_null c = true ∧
- (∀a,l1,x0,a0,l2,l3.
- ls = (l1@〈a0,false〉::〈x0,true〉::l2) →
- rs = (〈a,false〉::l3) →
- no_marks l1 →
- ((∃x. c = bit x ∧
- t2 = midtape STape (〈bit x,false〉::l1@〈a0,true〉::〈bit x,false〉::l2) 〈a,true〉 l3) ∨
- (c = null ∧
- t2 = midtape ? (〈null,false〉::l1@〈a0,true〉::〈x0,false〉::l2) 〈a,true〉 l3))).
-
-definition R_copy_step_false ≝
- λt1,t2.
- ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs →
- bit_or_null (\fst c) = false ∧ t2 = t1.
-
-lemma sem_copy_step :
- accRealize ? copy_step (inr … (inl … (inr … start_nop))) R_copy_step_true R_copy_step_false.
-#intape
-@(acc_sem_if_app … (sem_test_char ? (λc:STape.bit_or_null (\fst c))) …
- (sem_copy_step_subcase FSUnialpha (bit false) …
- (sem_copy_step_subcase FSUnialpha (bit true) … (sem_nocopy_subcase …)))
- (sem_nop …))
-[ #t1 #t2 #t3 whd in ⊢ (%→%→?); #H1 #H2 #ls #c #rs #Ht1 >Ht1 in H1;
- * * #c0 * whd in ⊢ (??%?→?); #Hc0 destruct (Hc0) #Hc #Ht3 % //
- #a #l1 #x0 #a0 #l2 #l3 #Hls #Hrs #Hl1marks >Hls in Ht3; >Hrs #Ht3
- cases (H2 … Ht3 ?)
- [ * #Hc' #Ht2 % %{false} % // <Hc' @Ht2
- | * #Hnotfalse whd in ⊢ (%→?); #Ht2 cases (Ht2 … Ht3 ?) -Ht2
- [ * #Hc' #Ht2 % %{true} % // <Hc' @Ht2
- | * #Hnottrue whd in ⊢ (%→?); #Ht2 cases (Ht2 … Ht3 ?) -Ht2
- [ * #Hc' #Ht2 %2 <Hc' % // @Ht2
- | * #Hnotnull @False_ind
- generalize in match Hnotnull;generalize in match Hnottrue;generalize in match Hnotfalse;
- cases c in Hc; normalize
- [ * [ #_ #_ * #Hfalse #_ | #_ * #Hfalse #_ #_ ]
- | #_ #_ #_ * #Hfalse
- |*: #Hfalse destruct (Hfalse) ] @Hfalse %
- | @Hl1marks ]
- | @Hl1marks ]
- | @Hl1marks ]
-| #t1 #t2 #t3 whd in ⊢ (%→%→?); #H1 #H2 #ls #c #rs #Ht1
- >Ht1 in H1; * #Hcur #Ht3 % // @Hcur % ]
-qed.
-
-(*
-1) il primo carattere è marcato
-2) l'ultimo carattere è l'unico che può essere null, gli altri sono bit
-3) il terminatore non è né bit, né null
-*)
-
-definition copy0 ≝ whileTM ? copy_step (inr … (inl … (inr … start_nop))).
-
-let rec merge_config (l1,l2:list STape) ≝
- match l1 with
- [ nil ⇒ nil ?
- | cons p1 l1' ⇒ match l2 with
- [ nil ⇒ nil ?
- | cons p2 l2' ⇒
- let 〈c1,b1〉 ≝ p1 in let 〈c2,b2〉 ≝ p2 in
- match c2 with
- [ null ⇒ p1
- | _ ⇒ p2 ] :: merge_config l1' l2' ] ].
-
-lemma merge_config_append :
- ∀l1,l2,l3,l4.|l1| = |l2| →
- merge_config (l1@l3) (l2@l4) = merge_config l1 l2@merge_config l3 l4.
-#l1 #l2 #l3 #l4 #Hlen @(list_ind2 … Hlen)
-[normalize //
-| #t1 #t2 * #c1 #b1 * #c2 #b2 #IH whd in ⊢ (??%%); >IH % ]
-qed.
-
-definition R_copy0 ≝ λt1,t2.
- ∀ls,c,c0,rs,l1,l3,l4.
- t1 = midtape STape (l3@l4@〈c0,true〉::ls) 〈c,true〉 (l1@rs) →
- no_marks l1 → no_marks (l3@l4) → |l1| = |l4| →
- ∀l1',bv.〈c,false〉::l1 = l1'@[〈comma,bv〉] → only_bits_or_nulls l1' →
- ∀l4',bg.l4@[〈c0,false〉] = 〈grid,bg〉::l4' → only_bits_or_nulls l4' →
- (c = comma ∧ t2 = t1) ∨
- (c ≠ comma ∧
- t2 = midtape ? (reverse ? l1'@l3@〈grid,true〉::
- merge_config l4' (reverse ? l1')@ls)
- 〈comma,true〉 rs).
-
-lemma inj_append_singleton_l1 :
- ∀A.∀l1,l2:list A.∀a1,a2.l1@[a1] = l2@[a2] → l1 = l2.
-#A #l1 #l2 #a1 #a2 #H lapply (eq_f … (reverse ?) … H)
->reverse_append >reverse_append normalize #H1 destruct
-lapply (eq_f … (reverse ?) … e0) >reverse_reverse >reverse_reverse //
-qed.
-
-lemma inj_append_singleton_l2 :
- ∀A.∀l1,l2:list A.∀a1,a2.l1@[a1] = l2@[a2] → a1 = a2.
-#A #l1 #l2 #a1 #a2 #H lapply (eq_f … (reverse ?) … H)
->reverse_append >reverse_append normalize #H1 destruct %
-qed.
-
-axiom daemon : ∀P:Prop.P.
-
-lemma list_cases2_full :
- ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:Prop.
- length ? l1 = length ? l2 →
- (l1 = [] → l2 = [] → P) →
- (∀hd1,hd2,tl1,tl2.l1 = hd1::tl1 → l2 = hd2::tl2 → P) → P.
-#T1 #T2 #l1 #l2 #P #Hl @(list_ind2 … Hl)
-[ #Pnil #Pcons @Pnil //
-| #tl1 #tl2 #hd1 #hd2 #IH1 #IH2 #Hp @Hp // ]
-qed.
-
-lemma wsem_copy0 : WRealize ? copy0 R_copy0.
-#intape #k #outc #Hloop
-lapply (sem_while … sem_copy_step intape k outc Hloop) [%] -Hloop
-* #ta * #Hstar @(star_ind_l ??????? Hstar)
-[ 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 % %
- [ generalize in match Hl1bits; -Hl1bits cases l1' in Hl1;
- [ normalize #Hl1 #c1 destruct (Hl1) %
- | * #c' #b' #l0 #Heq normalize in Heq; destruct (Heq)
- #Hl1bits lapply (Hl1bits 〈c',false〉 ?) [ @memb_hd ]
- >Hc #Hfalse destruct ]
- | @Houtc ]
-| #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
- cases (Htc … Htb) -Htc #Hcbitnull #Htc
- % [ % #Hc' >Hc' in Hcbitnull; normalize #Hfalse destruct (Hfalse) ]
- cut (|l1| = |reverse ? l4|) [>length_reverse @Hlen] #Hlen1
- @(list_cases2_full … Hlen1)
- [ (* case l1 = [] is discriminated because l1 contains at least comma *)
- #Hl1nil @False_ind >Hl1nil in Hl1; cases l1' normalize
- [ #Hl1 destruct normalize in Hcbitnull; destruct (Hcbitnull)
- | #p0 #l0 normalize #Hfalse destruct (Hfalse) cases l0 in e0;
- [ normalize #Hfalse1 destruct (Hfalse1)
- | #p0' #l0' normalize #Hfalse1 destruct (Hfalse1) ] ]
- | (* case c::l1 = c::a::l1'' *)
- * #a #ba * #a0 #ba0 #l1'' #l4'' #Hl1cons #Hl4cons
- lapply (eq_f ?? (reverse ?) ?? Hl4cons) >reverse_reverse >reverse_cons -Hl4cons #Hl4cons
- cut (ba = false)
- [ >Hl1cons in Hl1nomarks; #Hl1nomarks lapply (Hl1nomarks 〈a,ba〉 ?)
- [ @memb_hd | normalize // ] ] #Hba
- cut (ba0 = false)
- [ >Hl4cons in Hl3l4nomarks; #Hl3l4nomarks lapply (Hl3l4nomarks 〈a0,ba0〉 ?)
- [ @memb_append_l2 @memb_append_l2 @memb_hd | normalize // ] ] #Hba0
- >Hba0 in Hl4cons; >Hba in Hl1cons; -Hba0 -Hba #Hl1cons #Hl4cons
- >Hl4cons in Htc; >Hl1cons #Htc
- lapply (Htc a (l3@reverse ? l4'') c0 a0 ls (l1''@rs) ? (refl ??) ?)
- [ #x #Hx @Hl3l4nomarks >Hl4cons <associative_append
- @memb_append_l1 @Hx
- | >associative_append >associative_append %
- | -Htc
- cut (∃la.l1' = 〈c,false〉::la)
- [ >Hl1cons in Hl1; cases l1'
- [normalize #Hfalse destruct (Hfalse)
- | #p #la normalize #Hla destruct (Hla) @(ex_intro ?? la) % ] ]
- * #la #Hla
- cut (∃lb.l4' = lb@[〈c0,false〉])
- [ >Hl4cons in Hl4;
- @(list_elim_left … l4')
- [ #Heq lapply (eq_f … (length ?) … Heq)
- >length_append >length_append
- >commutative_plus normalize >commutative_plus normalize
- #Hfalse destruct
- | #a1 #tl #_ #Heq
- >(inj_append_singleton_l2 ? (reverse ? l4''@[〈a0,false〉]) (〈grid,bg〉::tl) 〈c0,false〉 a1 Heq)
- @ex_intro //
- ] ] * #lb #Hlb
- cut (|lb| = |reverse ? la|)
- [ >Hla in Hl1; >Hlb in Hl4; #Hl4 #Hl1
- >(?:l1 = la@[〈comma,bv〉]) in Hlen;
- [|normalize in Hl1; destruct (Hl1) %]
- >(?:l4 = 〈grid,bg〉::lb)
- [|@(inj_append_singleton_l1 ?? (〈grid,bg〉::lb) ?? Hl4) ]
- >length_append >commutative_plus >length_reverse
- normalize #Hlalb destruct (Hlalb) //
- ] #Hlen2
- *
- (* by hyp on the first iteration step,
- we consider whether c = bit x or c = null *)
- (* c = bit x *)
- [ * #x * #Hx #Htc
- lapply (Hind (〈bit x,false〉::ls) a a0 rs l1''
- (〈bit x,false〉::l3) (reverse ? l4'') ????)
- [ >Hl1cons in Hlen; >Hl4cons >length_append >commutative_plus
- normalize #Hlen destruct (Hlen) //
- | #x0 #Hx0 cases (orb_true_l … Hx0)
- [ #Hx0eq >(\P Hx0eq) %
- | -Hx0 #Hx0 @Hl3l4nomarks >Hl4cons
- <associative_append @memb_append_l1 // ]
- | #x0 #Hx0 @Hl1nomarks >Hl1cons @memb_cons //
- | >Htc >associative_append %
- | -Hind
- <Hl1cons <Hl4cons #Hind lapply (Hind la bv ?? lb bg ??)
- [ #x0 #Hx0 @Hl4bits >Hlb @memb_append_l1 //
- | >Hlb in Hl4; normalize in ⊢ (%→?); #Hl4
- @(inj_append_singleton_l1 ? l4 (〈grid,bg〉::lb) … Hl4)
- | #x0 #Hx0 @Hl1bits >Hla @memb_cons //
- | >Hla in Hl1; normalize in ⊢ (%→?); #Hl1
- destruct (Hl1) // ] -Hind
- (* by IH, we proceed by cases, whether a = comma
- (consequently several lists = []) or not *)
- *
- [ * #Ha #Houtc1
- cut (la = [] ∧ lb = [] ∧ l1'' = [] ∧ l4'' = [])
- [ @daemon ] * * * #Hla1 #Hlb1 #Hl1nil #Hl4nil
- >Hl1cons in Hl1; >Hla
- >Houtc1 >Htc #Hl1
- >Hl4cons in Hl4; >Hlb #Hl4
- >Hla1 >Hlb1 >Hl1nil >Hl4nil >Hx
- cut (a0 = grid) [ @daemon ] #Ha0 <Ha <Ha0
- normalize in ⊢ (??(??%?%)(??%?%)); >associative_append %
- | * #Ha #Houtc1 >Houtc1 @eq_f3 //
- >Hla >reverse_cons >associative_append @eq_f
- >Hx whd in ⊢ (??%?); @eq_f whd in ⊢ (???%); @eq_f @eq_f
- >Hlb >append_cons @eq_f2 // >(merge_config_append … Hlen2) %
- ]
- ]
- | (* c = null *)
- * #Hc #Htc
- lapply (Hind (〈c0,false〉::ls) a a0 rs l1'' (〈null,false〉::l3) (reverse ? l4'') ????)
- [ >Hl1cons in Hlen; >Hl4cons >length_append >commutative_plus normalize
- #Hlen destruct (Hlen) @e0
- | #x0 #Hx0 cases (memb_append STape ? [〈null,false〉] (l3@reverse ? l4'') … Hx0) -Hx0 #Hx0
- [ >(memb_single … Hx0) %
- | @Hl3l4nomarks cases (memb_append … Hx0) -Hx0 #Hx0
- [ @memb_append_l1 //
- | @memb_append_l2 >Hl4cons @memb_append_l1 // ]
- ]
- | >Hl1cons #x' #Hx0 @Hl1nomarks >Hl1cons @memb_cons //
- | >Htc @eq_f3 // >associative_append % ] -Hind <Hl1cons <Hl4cons #Hind
- lapply (Hind la bv ?? lb bg ??)
- [ #x0 #Hx0 @Hl4bits >Hlb @memb_append_l1 //
- | >Hlb in Hl4; normalize in ⊢ (%→?); #Hl4
- @(inj_append_singleton_l1 ? l4 (〈grid,bg〉::lb) … Hl4)
- | #x0 #Hx0 @Hl1bits >Hla @memb_cons //
- | >Hla in Hl1; normalize in ⊢ (%→?); #Hl1
- destruct (Hl1) // ] -Hind *
- (* by IH, we proceed by cases, whether a = comma
- (consequently several lists = []) or not *)
- [ * #Ha #Houtc1 >Hl1cons in Hl1; >Hla
- >Houtc1 >Htc #Hl1
- >Hl4cons in Hl4; >Hlb #Hl4
- cut (la = [] ∧ lb = [] ∧ l1'' = [] ∧ l4'' = [])
- [@daemon] * * * #Hla1 #Hlb1 #Hl1nil #Hl4nil
- >Hla1 >Hlb1 >Hl1nil >Hl4nil >Hc
- cut (a0 = grid) [ @daemon ] #Ha0 <Ha <Ha0
- normalize in ⊢ (??(??%?%)(??%?%)); >associative_append %
- | * #Ha #Houtc1 >Houtc1 @eq_f3 //
- >Hla >reverse_cons >associative_append @eq_f
- >Hc whd in ⊢ (??%?); @eq_f whd in ⊢ (???%); @eq_f @eq_f
- >Hlb >append_cons @eq_f2 // >(merge_config_append … Hlen2) %
- ]
- ]
-]]]
-qed.
-
-definition Pre_copy0 ≝ λt1.
- ∃ls,c,c0,rs,l1,l3,l4.
- t1 = midtape STape (l3@l4@〈c0,true〉::ls) 〈c,true〉 (l1@rs) ∧
- no_marks l1 ∧ no_marks (l3@l4) ∧ |l1| = |l4| ∧
- (∃l1',bv.〈c,false〉::l1 = l1'@[〈comma,bv〉] ∧ only_bits_or_nulls l1') ∧
- (∃l4',bg.l4@[〈c0,false〉] = 〈grid,bg〉::l4' ∧ only_bits_or_nulls l4').
-
-definition Pre_copy ≝ λt1.
- ∃ls,s0,s1,c0,c1,rs,l1,l3,l4.
- t1 = midtape STape (l3@〈grid,false〉::〈c0,false〉::l4@〈s0,true〉::ls) 〈s1,true〉 (l1@〈c1,false〉::〈comma,false〉::rs) ∧
- no_marks l1 ∧ no_marks l3 ∧ no_marks l4 ∧ |l1| = |l4| ∧
- only_bits (l4@[〈s0,false〉]) ∧ only_bits (〈s1,false〉::l1) ∧
- bit_or_null c0 = true ∧ bit_or_null c1 = true.
-
-lemma list_last: ∀A.∀l:list A.
- l = [ ] ∨ ∃a,l1. l = l1@[a].
-#A #l <(reverse_reverse ? l) cases (reverse A l)
- [%1 //
- |#a #l1 %2 @(ex_intro ?? a) @(ex_intro ?? (reverse ? l1)) //
- ]
-qed.
-
-lemma terminate_copy0 : ∀t.Pre_copy0 t → Terminate ? copy0 t.
-#t #HPre
-@(terminate_while_guarded ???
- Pre_copy0 …
- (acc_Realize_to_acc_GRealize ??? Pre_copy0 … sem_copy_step)
- … HPre) [%]
- [ -HPre -t #t1 #t2 * #ls * #c * #c0 * #rs * #l1 * #l3 * #l4
- * * * * * #Ht1 #Hl1nomarks #Hl3l4nomarks #Hlen
- * #l1' * #bv * #Hl1 #Hbitsnullsl1' * #l4' * #bg * #Hl4 #Hbitsnullsl4'
- #HR cases (HR … Ht1) -HR #Hbitnullc
- cut (∃d1,l1''.l1 = 〈d1,false〉::l1'')
- [ lapply Hl1nomarks cases l1 in Hl1;
- [ whd in ⊢ (???%→?); #Hl1
- lapply (append_l2_injective_r ? [] l1' [〈c,false〉] [〈comma,bv〉] (refl ??) Hl1)
- #Hc destruct (Hc) normalize in Hbitnullc; destruct (Hbitnullc)
- | * #d #bd #l1'' #_ #Hl1nomarks >(?:bd = false) [| @(Hl1nomarks 〈d,bd〉) @memb_hd ] /3/ ] ]
- * #d1 * #l1'' #Hl1''
- cut (∃d2,l4''.l4 = l4''@[〈d2,false〉])
- [ lapply Hl4 cases (list_last ? l4)
- [ #Hl4' >Hl4' in Hlen; >Hl1'' normalize in ⊢ (%→?); #HFalse destruct (HFalse)
- | * * #d2 #bd2 * #l4'' #Hl4'' >Hl4'' in Hl3l4nomarks; #Hl3l4nomarks #_
- <(?:bd2 = false) [| @(Hl3l4nomarks 〈d2,bd2〉) @memb_append_l2 @memb_append_l2 @memb_hd ]
- /3/ ] ]
- * #d2 * #l4'' #Hl4'' >Hl1'' >Hl4''
- cut (∃l1''',bv0.〈d1,false〉::l1''=l1'''@[〈comma,bv0〉]∧only_bits_or_nulls l1''')
- [ <Hl1'' cases (list_last ? l1) in Hl1'';
- [ #Hl1'' >Hl1''#HFalse destruct(HFalse)
- | * #a * #l1''' #Hl1''' >Hl1''' in Hl1; #Hl1 #_
- lapply (append_l2_injective_r ? (〈c,false〉::l1''') l1' [a] [〈comma,bv〉] (refl …) Hl1)
- #Ha destruct (Ha) %{l1'''} %{bv} % //
- #x #Hx @Hbitsnullsl1'
- lapply (append_l1_injective_r ? (〈c,false〉::l1''') l1' [〈comma,bv〉] [〈comma,bv〉] (refl …) Hl1)
- #H <H @memb_cons @Hx ] ]
- cut (∃l4''',bg0.l4''@[〈d2,false〉]=〈grid,bg0〉::l4'''∧only_bits_or_nulls l4''')
- [ <Hl4'' cases (list_last ? l4') in Hl4;
- [ #Hl4' >Hl4' >Hl4'' cases l4''
- [ normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
- | #y #yl normalize in ⊢ (%→?); #H1 destruct (H1) cases yl in e0;
- [ normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
- | #z #zl normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ] ]
- | * #a * #l4''' #Hl4''' >Hl4''' #Hl4
- lapply (append_l1_injective_r ? l4 (〈grid,bg〉::l4''') [〈c0,false〉] [a] (refl …) Hl4)
- -Hl4 #Hl4 >Hl4 %{l4'''} %{bg} % //
- #x #Hx @Hbitsnullsl4' >Hl4''' @memb_append_l1 @Hx ] ]
- #Hl4''' #Hl1'''
- #HR cases (HR d1 (l3@l4'') c0 d2 ls (l1''@rs) ? ? ?)
- [3: >associative_append >associative_append %
- |4: %
- |5: #x #Hx @Hl3l4nomarks cases (memb_append … Hx)
- [ @memb_append_l1
- | >Hl4'' -Hx #Hx @memb_append_l2 @memb_append_l1 @Hx ]
- | * #x1 * #Hc #Ht2 whd >Ht2
- %{(〈bit x1,false〉::ls)} %{d1} %{d2} %{rs} %{l1''} %{(〈bit x1,false〉::l3)} %{l4''}
- % [ % [ % [ % [ %
- [ >associative_append %
- | #x #Hx @Hl1nomarks >Hl1'' @memb_cons @Hx ]
- | #x #Hx cases (orb_true_l … Hx) -Hx #Hx
- [ >(\P Hx) %
- | @Hl3l4nomarks cases (memb_append … Hx)
- [ @memb_append_l1
- | -Hx #Hx >Hl4'' @memb_append_l2 @memb_append_l1 @Hx ] ] ]
- | >Hl1'' in Hlen; >Hl4'' >length_append >commutative_plus
- normalize in ⊢ (%→?); #Hlen destruct (Hlen) @e0 ]
- | @Hl1''' ]
- | @Hl4''' ]
- | * #Hc #Ht2 whd >Ht2
- %{(〈c0,false〉::ls)} %{d1} %{d2} %{rs} %{l1''} %{(〈null,false〉::l3)} %{l4''}
- % [ % [ % [ % [ %
- [ >associative_append %
- | #x #Hx @Hl1nomarks >Hl1'' @memb_cons @Hx ]
- | #x #Hx cases (orb_true_l … Hx) -Hx #Hx
- [ >(\P Hx) %
- | @Hl3l4nomarks cases (memb_append … Hx)
- [ @memb_append_l1
- | -Hx #Hx >Hl4'' @memb_append_l2 @memb_append_l1 @Hx ] ] ]
- | >Hl1'' in Hlen; >Hl4'' >length_append >commutative_plus
- normalize in ⊢ (%→?); #Hlen destruct (Hlen) @e0 ]
- | @Hl1''' ]
- | @Hl4''' ]
- ]
- |cases HPre -HPre #ls * #c * #c0 * #rs * #l1 * #l3 * #l4
- * * * * * #Ht #Hl1nomarks #Hl3l4nomarks #Hlen
- * #l1' * #bv * #Hl1 #Hbitsnullsl1'
- * #l4' * #bg * #Hl4 #Hbitsnullsl4'
- >Ht lapply Hbitsnullsl1' lapply Hl1 lapply l1' lapply Hl3l4nomarks
- lapply Hl1nomarks lapply l3 lapply c0 lapply c lapply ls
- -Hbitsnullsl1' -Hl1 -l1' -Hl3l4nomarks -Hl1nomarks -l3 -Hl4 -c0 -c -ls
- <(reverse_reverse ? l4) <(length_reverse ? l4) in Hlen; #Hlen
- @(list_ind2 … Hlen)
- [ #ls #c #c0 #l3 #_ #_ #l1' #Hl1 #Hbitsnullsl1' cases l1' in Hl1;
- [| #x * [| #x0 #xs ] normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
- normalize in ⊢ (%→?); #Hl1' destruct (Hl1') % #t1 whd in ⊢ (%→?);
- #HR cases (HR … (refl …)) whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse)
- | #xs #ys * #x #bx * #y #by #IH
- #ls #c #c0 #l3 #Hl1nomarks #Hl3l4nomarks #l1' #Hl1
- #Hbitsnullsl1' %
- #t1 whd in ⊢ (%→?); #HR cases (HR … (refl …)) -HR
- #Hbitnullc #HR
- lapply (Hl1nomarks 〈x,bx〉 (memb_hd …)) normalize in ⊢ (%→?); #Hbx destruct (Hbx)
- (*cut (∃d2,l4''.〈y,by〉::ys = l4''@[〈d2,false〉])
- [ cases (list_last ? ys) in Hl3l4nomarks;
- [ #Hl4' #Hl3l4nomarks >Hl4' >(?:by = false) [%{y} %{([])} %]
- @(Hl3l4nomarks 〈y,by〉) @memb_append_l2 @memb_hd
- | * * #d2 #bd2 * #l4'' #Hl4'' >Hl4'' #Hl3l4nomarks
- <(?:bd2 = false) [| @(Hl3l4nomarks 〈d2,bd2〉) @memb_append_l2 @memb_cons @memb_append_l2 @memb_hd ]
- %{d2} %{(〈y,by〉::l4'')} % ] ]
- * #a0 * #l4'' #Hl4'' >Hl4'' in HR; #HR *)
- cases (HR x (l3@reverse ? ys) c0 y ls (xs@rs) ? (refl …) ?)
- [3: >reverse_cons >associative_append >associative_append
- >(?:by = false) [|@(Hl3l4nomarks 〈y,by〉) @memb_append_l2 >reverse_cons @memb_append_l2 @memb_hd ] %
- |4: #x #Hx @Hl3l4nomarks >reverse_cons <associative_append @memb_append_l1 @Hx
- | * #x1 * #Hc #Ht1 >Ht1
- <(?: (〈bit x1,false〉::l3)@reverse (FSUnialpha×bool) ys@〈y,true〉::〈bit x1,false〉::ls =
- 〈bit x1,false〉::(l3@reverse (FSUnialpha×bool) ys)@〈y,true〉::〈bit x1,false〉::ls)
- [cut (∃l1''.〈x,false〉::xs = l1''@[〈comma,bv〉] ∧ only_bits_or_nulls l1'')
- [lapply Hbitsnullsl1' lapply Hl1 -Hl1 -Hbitsnullsl1' cases l1'
- [normalize in ⊢(%→?); #Hfalse destruct (Hfalse)
- |#x2 #l1'' normalize in ⊢ (%→?); #Heq #Hbitsnullsl1' destruct (Heq)
- %{l1''} % [@e0]
- #x #Hx @Hbitsnullsl1' @memb_cons @Hx ] ]
- * #l1'' * #Hl1'' #Hbitsnullsl1''
- @(IH (〈bit x1,false〉::ls) x y (〈bit x1,false〉::l3) ?? l1'' Hl1'' Hbitsnullsl1'')
- [#x #Hx @Hl1nomarks @memb_cons @Hx
- |#x #Hx cases (orb_true_l … Hx) -Hx #Hx
- [>(\P Hx) %
- |cases (memb_append … Hx) -Hx #Hx @Hl3l4nomarks
- [@memb_append_l1 @Hx
- |@memb_append_l2 >reverse_cons @memb_append_l1 @Hx ] ] ]
- |>associative_append % ]
- | * #Hc #Ht1 >Ht1
- <(?: (〈null,false〉::l3)@reverse (FSUnialpha×bool) ys@〈y,true〉::〈c0,false〉::ls =
- 〈null,false〉::(l3@reverse (FSUnialpha×bool) ys)@〈y,true〉::〈c0,false〉::ls)
- [cut (∃l1''.〈x,false〉::xs = l1''@[〈comma,bv〉] ∧ only_bits_or_nulls l1'')
- [lapply Hbitsnullsl1' lapply Hl1 -Hl1 -Hbitsnullsl1' cases l1'
- [normalize in ⊢(%→?); #Hfalse destruct (Hfalse)
- |#x2 #l1'' normalize in ⊢ (%→?); #Heq #Hbitsnullsl1' destruct (Heq)
- %{l1''} % [@e0]
- #x #Hx @Hbitsnullsl1' @memb_cons @Hx ] ]
- * #l1'' * #Hl1'' #Hbitsnullsl1''
- @(IH (〈c0,false〉::ls) x y (〈null,false〉::l3) ?? l1'' Hl1'' Hbitsnullsl1'')
- [#x #Hx @Hl1nomarks @memb_cons @Hx
- |#x #Hx cases (orb_true_l … Hx) -Hx #Hx
- [>(\P Hx) %
- |cases (memb_append … Hx) -Hx #Hx @Hl3l4nomarks
- [@memb_append_l1 @Hx
- |@memb_append_l2 >reverse_cons @memb_append_l1 @Hx ] ] ]
- |>associative_append % ]
-]]]
-qed.
-
-definition merge_char ≝ λc1,c2.
- match c2 with
- [ null ⇒ c1
- | _ ⇒ c2 ].
-
-lemma merge_cons :
- ∀c1,c2,conf1,conf2.
- merge_config (〈c1,false〉::conf1) (〈c2,false〉::conf2) =
- 〈merge_char c1 c2,false〉::merge_config conf1 conf2.
-#c1 #c2 #conf1 #conf2 normalize @eq_f2 //
-cases c2 /2/
-qed.
-
-lemma merge_bits : ∀l1,l2.|l1| = |l2| → only_bits l2 → merge_config l1 l2 = l2.
-#l1 #l2 #Hlen @(list_ind2 … Hlen) //
-#tl1 #tl2 #hd1 #hd2 #IH
->(eq_pair_fst_snd … hd1) >(eq_pair_fst_snd … hd2) #Hbits
-change with (cons ???) in ⊢ (??%?); @eq_f2
-[ cases (\fst hd2) in Hbits;
- [ #b #_ %
- |*: #Hfalse lapply (Hfalse … (memb_hd …)) normalize #Hfalse1 destruct (Hfalse1) ]
-| @IH #x #Hx @Hbits @memb_cons // ]
-qed.
-
-lemma merge_config_c_nil :
- ∀c.merge_config c [] = [].
-#c cases c normalize //
-qed.
-
-lemma reverse_merge_config :
- ∀c1,c2.|c1| = |c2| → reverse ? (merge_config c1 c2) =
- merge_config (reverse ? c1) (reverse ? c2).
-#c1 #c2 <(length_reverse ? c1) <(length_reverse ? c2) #Hlen
-<(reverse_reverse ? c1) in ⊢ (??%?); <(reverse_reverse ? c2) in ⊢ (??%?);
-generalize in match Hlen; @(list_ind2 … Hlen) -Hlen //
-#tl1 #tl2 #hd1 #hd2 #IH whd in ⊢ (??%%→?); #Hlen destruct (Hlen) -Hlen
-<(length_reverse ? tl1) in e0; <(length_reverse ? tl2) #Hlen
->reverse_cons >reverse_cons >(merge_config_append ???? Hlen)
->reverse_append >(eq_pair_fst_snd ?? hd1) >(eq_pair_fst_snd ?? hd2)
-whd in ⊢ (??%%); @eq_f2 // @IH //
-qed.
-
-definition copy
-≝
- seq STape copy0 (seq ? (move_l …) (seq ? (adv_to_mark_l … (is_marked ?))
- (seq ? (clear_mark …) (seq ? (adv_to_mark_r … (is_marked ?)) (clear_mark …))))).
-
-(*
- s0, s1 = caratteri di testa dello stato
- c0 = carattere corrente del nastro oggetto
- c1 = carattere in scrittura sul nastro oggetto
-
- questa dimostrazione sfrutta il fatto che
- merge_config (l0@[c0]) (l1@[c1]) = l1@[merge_char c0 c1]
- se l0 e l1 non contengono null
-*)
-
-definition R_copy ≝ λt1,t2.
- ∀ls,s0,s1,c0,c1,rs,l1,l3,l4.
- t1 = midtape STape (l3@〈grid,false〉::〈c0,false〉::l4@〈s0,true〉::ls) 〈s1,true〉 (l1@〈c1,false〉::〈comma,false〉::rs) →
- no_marks l1 → no_marks l3 → no_marks l4 → |l1| = |l4| →
- only_bits (l4@[〈s0,false〉]) → only_bits (〈s1,false〉::l1) →
- bit_or_null c0 = true → bit_or_null c1 = true →
- t2 = midtape STape (〈c1,false〉::reverse ? l1@〈s1,false〉::l3@〈grid,false〉::
- 〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls)
- 〈comma,false〉 rs.
-
-lemma sem_copy0 : GRealize ? copy0 Pre_copy0 R_copy0.
-@WRealize_to_GRealize
-[ @terminate_copy0
-| @wsem_copy0 ]
-qed.
-
-definition option_cons ≝ λA.λa:option A.λl.
- match a with
- [ None ⇒ l
- | Some a' ⇒ a'::l ].
-
-lemma sem_copy : GRealize ? copy Pre_copy R_copy.
-@(GRealize_to_GRealize_2 ?? Pre_copy0 ? R_copy) //
-(* preamble: Pre_copy0 implies Pre_copy *)
-[ #t1 * #ls * #s0 * #s1 * #c0 * #c1 * #rs * #l1 * #l3 * #l4 * * * * * * * *
- #Ht1 #Hl1nomarks #Hl3nomarks #Hl4nomarks #Hlen #Hbitsl4 #Hbitsl1
- #Hbitnullc0 #Hbitnullc1 whd
- %{ls} %{s1} %{s0} %{rs} %{(l1@[〈c1,false〉;〈comma,false〉])}
- %{l3} %{(〈grid,false〉::〈c0,false〉::l4)}
- % [ % [ % [ % [ %
- [ >Ht1 -Ht1 @eq_f >associative_append %
- | #x #Hx cases (memb_append … Hx) -Hx #Hx
- [ @(Hl1nomarks ? Hx)
- | cases (orb_true_l … Hx) -Hx #Hx
- [ >(\P Hx) %
- | >(memb_single … Hx) % ] ] ]
- | #x #Hx cases (memb_append … Hx) -Hx #Hx
- [ @(Hl3nomarks ? Hx)
- | cases (orb_true_l … Hx) -Hx #Hx
- [ >(\P Hx) %
- | cases (orb_true_l … Hx) -Hx #Hx
- [ >(\P Hx) %
- | @(Hl4nomarks ? Hx) ]]]]
- |>length_append >Hlen >commutative_plus % ]
- | %{(〈s1,false〉::l1@[〈c1,false〉])} %{false} %
- [ @eq_f >associative_append %
- | #x #Hx cases (orb_true_l … Hx) -Hx #Hx
- [ >(\P Hx) @is_bit_to_bit_or_null @(Hbitsl1 〈s1,false〉) @memb_hd
- | cases (memb_append … Hx) -Hx #Hx
- [ @is_bit_to_bit_or_null @(Hbitsl1 〈\fst x,\snd x〉) @memb_cons <eq_pair_fst_snd @Hx
- | >(memb_single… Hx) // ]
- ]
- ] ]
- | %{(〈c0,false〉::l4@[〈s0,false〉])} %{false} %
- [ %
- | #x #Hx cases (orb_true_l … Hx) -Hx #Hx
- [ >(\P Hx) //
- | cases (memb_append … Hx) -Hx #Hx
- [ @is_bit_to_bit_or_null @(Hbitsl4 〈\fst x,\snd x〉) @memb_append_l1 <eq_pair_fst_snd @Hx
- | >(memb_single… Hx) @is_bit_to_bit_or_null @(Hbitsl4 〈s0,false〉) @memb_append_l2 @memb_hd ]
- ]
-] ] ]
-(*whd in ⊢ (%→%)
-change with (?·?) in match copy;*)
-@(sem_seq_app_guarded ???? (λx.True) ??? sem_copy0)
-[2:@(sem_seq_app_guarded ???????? (Realize_to_GRealize … (sem_move_l (FinProd FSUnialpha FinBool))))
- [@(λx.True)
- |4://
- |5:@sub_reflexive
- |3:@(sem_seq_app_guarded ???????? (Realize_to_GRealize … (sem_adv_to_mark_l … (is_marked FSUnialpha))))
- [@(λx.True)
- |4://
- |5:@sub_reflexive
- |3:@(sem_seq_app_guarded ???????? (Realize_to_GRealize … (sem_clear_mark ?)))
- [@(λx.True)
- |4://
- |5:@sub_reflexive
- |3:@(sem_seq_app_guarded ???????? (Realize_to_GRealize … (sem_adv_to_mark_r ? (is_marked ?))))
- [@(λx.True)
- |4://
- |5:@sub_reflexive
- |3:@(Realize_to_GRealize … (sem_clear_mark ?)) ] ] ] ]
-|3: //
-|1:skip]
-#intape #outtape #HR
-#ls #s0 #s1 #c0 #c1 #rs #l1 #l2 #l3 #Hintape #Hl1marks #Hl2marks #Hl3marks #Hlen
-#Hbits1 #Hbits2 #Hc0bits #Hc1bits
-cases HR -HR #ta * whd in ⊢ (%→?); #Hta
-cut (ta = midtape STape (〈c1,false〉::reverse ? l1@〈s1,false〉::l2@〈grid,true〉::
- 〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls)
- 〈comma,true〉 rs)
-[lapply (Hta ls s1 s0 rs (l1@[〈c1,false〉;〈comma,false〉]) l2 (〈grid,false〉::〈c0,false〉::l3) ?)
- [>associative_append in ⊢ (???(????%)); normalize in ⊢ (???(??%%%)); @Hintape ]
- -Hta #Hta cases (Hta ??? (〈s1,false〉::l1@[〈c1,false〉]) false ? ? ?? (refl ??) ?)
- [3: #x #Hx cases (memb_append … Hx) -Hx #Hx
- [ @Hl1marks //
- | cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) % | >(memb_single … Hx) % ]]
- |4: #x #Hx cases (memb_append … Hx) -Hx #Hx
- [ @Hl2marks //
- | cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) % | cases (orb_true_l … Hx) [-Hx #Hx >(\P Hx) % | @Hl3marks ] ] ]
- |5: >length_append normalize >Hlen >commutative_plus %
- |6: normalize >associative_append %
- |7: #x #Hx cases (memb_append ?? (〈s1,false〉::l1) … Hx) -Hx #Hx
- [ whd in ⊢ (??%?); >(Hbits2 … Hx) %
- | >(memb_single … Hx) // ]
- |8: #x #Hx cases (memb_append … Hx) -Hx #Hx
- [ cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) // | whd in ⊢ (??%?); >Hbits1 // @memb_append_l1 // ]
- | >(memb_single … Hx) whd in ⊢ (??%?); >(Hbits1 〈s0,false〉) // @memb_append_l2 @memb_hd ]
- | * #Hs1 @False_ind >Hs1 in Hbits2; #Hbits2 lapply (Hbits2 〈comma,false〉 ?) //
- normalize #Hfalse destruct (Hfalse)
- | * #Hs1 #Ht2 >Ht2 >reverse_cons >reverse_append >reverse_single @eq_f3 //
- >merge_cons >merge_bits
- [2: #x #Hx @Hbits2 cases (memb_append STape ? (reverse ? l1) ? Hx) -Hx #Hx
- [@daemon | >(memb_single … Hx) @memb_hd ]
- |3: >length_append >length_append >length_reverse >Hlen % ]
- normalize >associative_append normalize >associative_append %
- ]
-] -Hta #Hta * #tb * * #_ #Htb
-lapply (Htb … Hta) -Htb #Htb change with (midtape ????) in Htb:(???%);
-* #tc * * #_ #Htc
-cases (Htc … Htb)
-(* [ * #Hfalse normalize in Hfalse; destruct (Hfalse) ] *)
-#_ #Htc cases (Htc (refl ??)) -Htc #Htc #_
-lapply (Htc (reverse ? l1@〈s1,false〉::l2) 〈grid,true〉
- (〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls)???)
-[ #x #Hx cases (memb_append … Hx) -Hx #Hx
- [ @Hl1marks @daemon
- | cases (orb_true_l … Hx) -Hx #Hx
- [ >(\P Hx) % | @(Hl2marks … Hx) ] ]
-| %
-| whd in ⊢ (??%?); >associative_append % ] -Htc #Htc
-* #td * whd in ⊢ (%→?); * #_ #Htd lapply (Htd … Htc) -Htd #Htd
-* #te * whd in ⊢ (%→?); * #_ #Hte cases (Hte … Htd) -Hte -Htd
-[ * #Hfalse normalize in Hfalse; destruct (Hfalse) ]
-* * #_ #Hte #_
-lapply (Hte (reverse ? (reverse ? l1@〈s1,false〉::l2)@[〈c1,false〉])
- 〈comma,true〉 rs ? (refl ??) ?) -Hte
-[ >reverse_append >reverse_cons >reverse_reverse #x #Hx
- cases (memb_append … Hx) -Hx #Hx
- [ cases (memb_append … Hx) -Hx #Hx
- [ cases (memb_append … Hx) -Hx #Hx
- [ @daemon
- | lapply (memb_single … Hx) -Hx #Hx >Hx % ]
- | @(Hl1marks … Hx) ]
- | lapply (memb_single … Hx) -Hx #Hx >Hx % ]
-| >reverse_append >reverse_reverse >reverse_cons
- >associative_append >associative_append >associative_append
- >associative_append >associative_append % ]
-#Hte * #_ #Houtc lapply (Houtc … Hte) -Houtc #Houtc >Houtc
-@eq_f3 //
->reverse_append >reverse_append >reverse_single >reverse_cons
->reverse_append >reverse_append >reverse_reverse >reverse_reverse
->reverse_single >associative_append >associative_append %
-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_____________________________________________________________*)
-
-
-(* COMPARE BIT
-
-*)
-
-include "turing/if_machine.ma".
-include "turing/basic_machines.ma".
-include "turing/universal/alphabet.ma".
-
-(* ADVANCE TO MARK (right)
-
- sposta la testina a destra fino a raggiungere il primo carattere marcato
-
-*)
-
-(* 0, a ≠ mark _ ⇒ 0, R
-0, a = mark _ ⇒ 1, N *)
-
-definition atm_states ≝ initN 3.
-
-definition atm0 : atm_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
-definition atm1 : atm_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
-definition atm2 : atm_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
-
-definition atmr_step ≝
- λalpha:FinSet.λtest:alpha→bool.
- mk_TM alpha atm_states
- (λp.let 〈q,a〉 ≝ p in
- match a with
- [ None ⇒ 〈atm1, None ?〉
- | Some a' ⇒
- match test a' with
- [ true ⇒ 〈atm1,None ?〉
- | false ⇒ 〈atm2,Some ? 〈a',R〉〉 ]])
- atm0 (λx.notb (x == atm0)).
-
-definition Ratmr_step_true ≝
- λalpha,test,t1,t2.
- ∃ls,a,rs.
- t1 = midtape alpha ls a rs ∧ test a = false ∧
- t2 = mk_tape alpha (a::ls) (option_hd ? rs) (tail ? rs).
-
-definition Ratmr_step_false ≝
- λalpha,test,t1,t2.
- t1 = t2 ∧
- (current alpha t1 = None ? ∨
- (∃a.current ? t1 = Some ? a ∧ test a = true)).
-
-lemma atmr_q0_q1 :
- ∀alpha,test,ls,a0,rs. test a0 = true →
- step alpha (atmr_step alpha test)
- (mk_config ?? atm0 (midtape … ls a0 rs)) =
- mk_config alpha (states ? (atmr_step alpha test)) atm1
- (midtape … ls a0 rs).
-#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?);
-whd in match (trans … 〈?,?〉); >Htest %
-qed.
-
-lemma atmr_q0_q2 :
- ∀alpha,test,ls,a0,rs. test a0 = false →
- step alpha (atmr_step alpha test)
- (mk_config ?? atm0 (midtape … ls a0 rs)) =
- mk_config alpha (states ? (atmr_step alpha test)) atm2
- (mk_tape … (a0::ls) (option_hd ? rs) (tail ? rs)).
-#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?);
-whd in match (trans … 〈?,?〉); >Htest cases ts //
-qed.
-
-lemma sem_atmr_step :
- ∀alpha,test.
- accRealize alpha (atmr_step alpha test)
- atm2 (Ratmr_step_true alpha test) (Ratmr_step_false alpha test).
-#alpha #test *
-[ @(ex_intro ?? 2)
- @(ex_intro ?? (mk_config ?? atm1 (niltape ?))) %
- [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ]
-| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (leftof ? a al)))
- % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ]
-| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (rightof ? a al)))
- % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ]
-| #ls #c #rs @(ex_intro ?? 2)
- cases (true_or_false (test c)) #Htest
- [ @(ex_intro ?? (mk_config ?? atm1 ?))
- [| %
- [ %
- [ whd in ⊢ (??%?); >atmr_q0_q1 //
- | whd in ⊢ ((??%%)→?); #Hfalse destruct ]
- | #_ % // %2 @(ex_intro ?? c) % // ]
- ]
- | @(ex_intro ?? (mk_config ?? atm2 (mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs))))
- %
- [ %
- [ whd in ⊢ (??%?); >atmr_q0_q2 //
- | #_ @(ex_intro ?? ls) @(ex_intro ?? c) @(ex_intro ?? rs)
- % // % //
- ]
- | #Hfalse @False_ind @(absurd ?? Hfalse) %
- ]
- ]
-]
-qed.
-
-lemma dec_test: ∀alpha,rs,test.
- decidable (∀c.memb alpha c rs = true → test c = false).
-#alpha #rs #test elim rs
- [%1 #n normalize #H destruct
- |#a #tl cases (true_or_false (test a)) #Ha
- [#_ %2 % #Hall @(absurd ?? not_eq_true_false) <Ha
- @Hall @memb_hd
- |* [#Hall %1 #c #memc cases (orb_true_l … memc)
- [#eqca >(\P eqca) @Ha |@Hall]
- |#Hnall %2 @(not_to_not … Hnall) #Hall #c #memc @Hall @memb_cons //
- ]
- qed.
-
-definition R_adv_to_mark_r ≝ λalpha,test,t1,t2.
- (current ? t1 = None ? → t1 = t2) ∧
- ∀ls,c,rs.
- (t1 = midtape alpha ls c rs →
- ((test c = true ∧ t2 = t1) ∨
- (test c = false ∧
- (∀rs1,b,rs2. rs = rs1@b::rs2 →
- test b = true → (∀x.memb ? x rs1 = true → test x = false) →
- t2 = midtape ? (reverse ? rs1@c::ls) b rs2) ∧
- ((∀x.memb ? x rs = true → test x = false) →
- ∀a,l.reverse ? (c::rs) = a::l →
- t2 = rightof alpha a (l@ls))))).
-
-definition adv_to_mark_r ≝
- λalpha,test.whileTM alpha (atmr_step alpha test) atm2.
-
-lemma wsem_adv_to_mark_r :
- ∀alpha,test.
- WRealize alpha (adv_to_mark_r alpha test) (R_adv_to_mark_r alpha test).
-#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)
-[ * #Htapea *
- [ #H1 %
- [#_ @Htapea
- |#ls #c #rs #H2 >H2 in H1; whd in ⊢ (??%? → ?);
- #Hfalse destruct (Hfalse)
- ]
- | * #a * #Ha #Htest %
- [ >Ha #H destruct (H);
- | #ls #c #rs #H2 %
- >H2 in Ha; whd in ⊢ (??%? → ?); #Heq destruct (Heq) % //
- <Htapea //
- ]
- ]
-| #tapeb #tapec #Hleft #Hright #IH #HRfalse
- lapply (IH HRfalse) -IH #IH %
- [cases Hleft #ls * #a * #rs * * #Htapea #_ #_ >Htapea
- whd in ⊢((??%?)→?); #H destruct (H);
- |#ls #c #rs #Htapea %2
- cases Hleft #ls0 * #a0 * #rs0 * * #Htapea' #Htest #Htapeb
- >Htapea' in Htapea; #Htapea destruct (Htapea) % [ % // ]
- [*
- [ #b #rs2 #Hrs >Hrs in Htapeb; #Htapeb #Htestb #_
- cases (proj2 ?? IH … Htapeb)
- [ * #_ #Houtc >Houtc >Htapeb %
- | * * >Htestb #Hfalse destruct (Hfalse) ]
- | #r1 #rs1 #b #rs2 #Hrs >Hrs in Htapeb; #Htapeb #Htestb #Hmemb
- cases (proj2 ?? IH … Htapeb)
- [ * #Hfalse >(Hmemb …) in Hfalse;
- [ #Hft destruct (Hft)
- | @memb_hd ]
- | * * #Htestr1 #H1 #_ >reverse_cons >associative_append
- @H1 // #x #Hx @Hmemb @memb_cons //
- ]
- ]
- |cases rs in Htapeb; normalize in ⊢ (%→?);
- [#Htapeb #_ #a0 #l whd in ⊢ ((??%?)→?); #Hrev destruct (Hrev)
- >Htapeb in IH; #IH cases (proj1 ?? IH … (refl …)) //
- |#r1 #rs1 #Htapeb #Hmemb
- cases (proj2 ?? IH … Htapeb) [ * >Hmemb [ #Hfalse destruct(Hfalse) ] @memb_hd ]
- * #_ #H1 #a #l <(reverse_reverse … l) cases (reverse … l)
- [#H cut (c::r1::rs1 = [a])
- [<(reverse_reverse … (c::r1::rs1)) >H //]
- #Hrev destruct (Hrev)
- |#a1 #l2 >reverse_cons >reverse_cons >reverse_cons
- #Hrev cut ([c] = [a1])
- [@(append_l2_injective_r ?? (a::reverse … l2) … Hrev) //]
- #Ha <Ha >associative_append @H1
- [#x #membx @Hmemb @memb_cons @membx
- |<(append_l1_injective_r ?? (a::reverse … l2) … Hrev) //
- ]
-qed.
-
-lemma terminate_adv_to_mark_r :
- ∀alpha,test.
- ∀t.Terminate alpha (adv_to_mark_r alpha test) t.
-#alpha #test #t
-@(terminate_while … (sem_atmr_step alpha test))
- [ %
- | cases t
- [ % #t1 * #ls0 * #c0 * #rs0 * * #Hfalse destruct (Hfalse)
- |2,3: #a0 #al0 % #t1 * #ls0 * #c0 * #rs0 * * #Hfalse destruct (Hfalse)
- | #ls #c #rs generalize in match c; -c generalize in match ls; -ls
- elim rs
- [#ls #c % #t1 * #ls0 * #c0 * #rs0 * *
- #H1 destruct (H1) #Hc0 #Ht1 normalize in Ht1;
- % #t2 * #ls1 * #c1 * #rs1 * * >Ht1
- normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
- | #r0 #rs0 #IH #ls #c % #t1 * #ls0 * #c0 * #rs0 * *
- #H1 destruct (H1) #Hc0 #Ht1 normalize in Ht1;
- >Ht1 @IH
- ]
- ]
- ]
-qed.
-
-lemma sem_adv_to_mark_r :
- ∀alpha,test.
- Realize alpha (adv_to_mark_r alpha test) (R_adv_to_mark_r alpha test).
-/2/
-qed.
-
-(* MARK machine
-
- marks the current character
- *)
-
-definition mark_states ≝ initN 2.
-
-definition ms0 : mark_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)).
-definition ms1 : mark_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)).
-
-definition mark ≝
- λalpha:FinSet.mk_TM (FinProd … alpha FinBool) mark_states
- (λp.let 〈q,a〉 ≝ p in
- match a with
- [ None ⇒ 〈ms1,None ?〉
- | Some a' ⇒ match (pi1 … q) with
- [ O ⇒ let 〈a'',b〉 ≝ a' in 〈ms1,Some ? 〈〈a'',true〉,N〉〉
- | S q ⇒ 〈ms1,None ?〉 ] ])
- ms0 (λq.q == ms1).
-
-definition R_mark ≝ λalpha,t1,t2.
- (∀ls,c,b,rs.
- t1 = midtape (FinProd … alpha FinBool) ls 〈c,b〉 rs →
- t2 = midtape ? ls 〈c,true〉 rs) ∧
- (current ? t1 = None ? → t2 = t1).
-
-lemma sem_mark :
- ∀alpha.Realize ? (mark alpha) (R_mark alpha).
-#alpha #intape @(ex_intro ?? 2) cases intape
-[ @ex_intro
- [| % [ % | % [#ls #c #b #rs #Hfalse destruct | // ]]]
-|#a #al @ex_intro
- [| % [ % | % [#ls #c #b #rs #Hfalse destruct | // ]]]
-|#a #al @ex_intro
- [| % [ % | % [#ls #c #b #rs #Hfalse destruct ] // ]]
-| #ls * #c #b #rs
- @ex_intro [| % [ % | %
- [#ls0 #c0 #b0 #rs0 #H1 destruct (H1) %
- | whd in ⊢ ((??%?)→?); #H1 destruct (H1)]]]
-qed.
-
-
-(* MOVE RIGHT AND MARK machine
-
- marks the first character on the right
-
- (could be rewritten using (mark; move_right))
- *)
-
-definition mrm_states ≝ initN 3.
-
-definition mrm0 : mrm_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
-definition mrm1 : mrm_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
-definition mrm2 : mrm_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
-
-definition move_right_and_mark ≝
- λalpha:FinSet.mk_TM (FinProd … alpha FinBool) mrm_states
- (λp.let 〈q,a〉 ≝ p in
- match a with
- [ None ⇒ 〈mrm2,None ?〉
- | Some a' ⇒ match pi1 … q with
- [ O ⇒ 〈mrm1,Some ? 〈a',R〉〉
- | S q ⇒ match q with
- [ O ⇒ let 〈a'',b〉 ≝ a' in
- 〈mrm2,Some ? 〈〈a'',true〉,N〉〉
- | S _ ⇒ 〈mrm2,None ?〉 ] ] ])
- mrm0 (λq.q == mrm2).
-
-definition R_move_right_and_mark ≝ λalpha,t1,t2.
- ∀ls,c,d,b,rs.
- t1 = midtape (FinProd … alpha FinBool) ls c (〈d,b〉::rs) →
- t2 = midtape ? (c::ls) 〈d,true〉 rs.
-
-lemma sem_move_right_and_mark :
- ∀alpha.Realize ? (move_right_and_mark alpha) (R_move_right_and_mark alpha).
-#alpha #intape @(ex_intro ?? 3) cases intape
-[ @ex_intro
- [| % [ % | #ls #c #d #b #rs #Hfalse destruct ] ]
-|#a #al @ex_intro
- [| % [ % | #ls #c #d #b #rs #Hfalse destruct ] ]
-|#a #al @ex_intro
- [| % [ % | #ls #c #d #b #rs #Hfalse destruct ] ]
-| #ls #c *
- [ @ex_intro [| % [ % | #ls0 #c0 #d0 #b0 #rs0 #Hfalse destruct ] ]
- | * #d #b #rs @ex_intro
- [| % [ % | #ls0 #c0 #d0 #b0 #rs0 #H1 destruct (H1) % ] ] ] ]
-qed.
-
-(* CLEAR MARK machine
-
- clears the mark in the current character
- *)
-
-definition clear_mark_states ≝ initN 3.
-
-definition clear0 : clear_mark_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
-definition clear1 : clear_mark_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
-definition claer2 : clear_mark_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
-
-definition clear_mark ≝
- λalpha:FinSet.mk_TM (FinProd … alpha FinBool) clear_mark_states
- (λp.let 〈q,a〉 ≝ p in
- match a with
- [ None ⇒ 〈clear1,None ?〉
- | Some a' ⇒ match pi1 … q with
- [ O ⇒ let 〈a'',b〉 ≝ a' in 〈clear1,Some ? 〈〈a'',false〉,N〉〉
- | S q ⇒ 〈clear1,None ?〉 ] ])
- clear0 (λq.q == clear1).
-
-definition R_clear_mark ≝ λalpha,t1,t2.
- (current ? t1 = None ? → t1 = t2) ∧
- ∀ls,c,b,rs.
- t1 = midtape (FinProd … alpha FinBool) ls 〈c,b〉 rs →
- t2 = midtape ? ls 〈c,false〉 rs.
-
-lemma sem_clear_mark :
- ∀alpha.Realize ? (clear_mark alpha) (R_clear_mark alpha).
-#alpha #intape @(ex_intro ?? 2) cases intape
-[ @ex_intro
- [| % [ % | % [#_ %|#ls #c #b #rs #Hfalse destruct ]]]
-|#a #al @ex_intro
- [| % [ % | % [#_ %|#ls #c #b #rs #Hfalse destruct ]]]
-|#a #al @ex_intro
- [| % [ % | % [#_ %|#ls #c #b #rs #Hfalse destruct ]]]
-| #ls * #c #b #rs
- @ex_intro [| % [ % | %
- [whd in ⊢ (??%?→?); #H destruct| #ls0 #c0 #b0 #rs0 #H1 destruct (H1) % ]]]]
-qed.
-
-(* ADVANCE MARK RIGHT machine
-
- clears mark on current char,
- moves right, and marks new current char
-
-*)
-
-definition adv_mark_r ≝
- λalpha:FinSet.
- clear_mark alpha · move_r ? · mark alpha.
-
-definition R_adv_mark_r ≝ λalpha,t1,t2.
- (∀ls,c.
- (∀d,b,rs.
- t1 = midtape (FinProd … alpha FinBool) ls 〈c,true〉 (〈d,b〉::rs) →
- t2 = midtape ? (〈c,false〉::ls) 〈d,true〉 rs) ∧
- (t1 = midtape (FinProd … alpha FinBool) ls 〈c,true〉 [ ] →
- t2 = rightof ? 〈c,false〉 ls)) ∧
- (current ? t1 = None ? → t1 = t2).
-
-lemma sem_adv_mark_r :
- ∀alpha.Realize ? (adv_mark_r alpha) (R_adv_mark_r alpha).
-#alpha
-@(sem_seq_app … (sem_clear_mark …)
- (sem_seq ????? (sem_move_r ?) (sem_mark alpha)) …)
-#intape #outtape whd in ⊢ (%→?); * #ta *
-whd in ⊢ (%→?); #Hs1 whd in ⊢ (%→?); * #tb * #Hs2 whd in ⊢ (%→?); #Hs3 %
- [#ls #c %
- [#d #b #rs #Hintape @(proj1 … Hs3 ?? b ?)
- @(proj2 … Hs2 ls 〈c,false〉 (〈d,b〉::rs))
- @(proj2 ?? Hs1 … Hintape)
- |#Hintape lapply (proj2 ?? Hs1 … Hintape) #Hta lapply (proj2 … Hs2 … Hta)
- whd in ⊢ ((???%)→?); #Htb <Htb @(proj2 … Hs3) >Htb //
- ]
- |#Hcur lapply(proj1 ?? Hs1 … Hcur) #Hta >Hta >Hta in Hcur; #Hcur
- lapply (proj1 ?? Hs2 … Hcur) #Htb >Htb >Htb in Hcur; #Hcur
- @sym_eq @(proj2 ?? Hs3) @Hcur
- ]
-qed.
-
-(* ADVANCE TO MARK (left)
-
-axiomatized
-
-*)
-definition atml_step ≝
- λalpha:FinSet.λtest:alpha→bool.
- mk_TM alpha atm_states
- (λp.let 〈q,a〉 ≝ p in
- match a with
- [ None ⇒ 〈atm1, None ?〉
- | Some a' ⇒
- match test a' with
- [ true ⇒ 〈atm1,None ?〉
- | false ⇒ 〈atm2,Some ? 〈a',L〉〉 ]])
- atm0 (λx.notb (x == atm0)).
-
-definition Ratml_step_true ≝
- λalpha,test,t1,t2.
- ∃ls,a,rs.
- t1 = midtape alpha ls a rs ∧ test a = false ∧
- t2 = mk_tape alpha (tail ? ls) (option_hd ? ls) (a :: rs).
-
-definition Ratml_step_false ≝
- λalpha,test,t1,t2.
- t1 = t2 ∧
- (current alpha t1 = None ? ∨
- (∃a.current ? t1 = Some ? a ∧ test a = true)).
-
-lemma atml_q0_q1 :
- ∀alpha,test,ls,a0,rs. test a0 = true →
- step alpha (atml_step alpha test)
- (mk_config ?? atm0 (midtape … ls a0 rs)) =
- mk_config alpha (states ? (atml_step alpha test)) atm1
- (midtape … ls a0 rs).
-#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?);
-whd in match (trans … 〈?,?〉); >Htest %
-qed.
-
-lemma atml_q0_q2 :
- ∀alpha,test,ls,a0,rs. test a0 = false →
- step alpha (atml_step alpha test)
- (mk_config ?? atm0 (midtape … ls a0 rs)) =
- mk_config alpha (states ? (atml_step alpha test)) atm2
- (mk_tape … (tail ? ls) (option_hd ? ls) (a0 :: rs)).
-#alpha #test #ls #a0 #rs #Htest whd in ⊢ (??%?);
-whd in match (trans … 〈?,?〉); >Htest cases ls //
-qed.
-
-lemma sem_atml_step :
- ∀alpha,test.
- accRealize alpha (atml_step alpha test)
- atm2 (Ratml_step_true alpha test) (Ratml_step_false alpha test).
-#alpha #test *
-[ @(ex_intro ?? 2)
- @(ex_intro ?? (mk_config ?? atm1 (niltape ?))) %
- [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ]
-| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (leftof ? a al)))
- % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ]
-| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (rightof ? a al)))
- % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ]
-| #ls #c #rs @(ex_intro ?? 2)
- cases (true_or_false (test c)) #Htest
- [ @(ex_intro ?? (mk_config ?? atm1 ?))
- [| %
- [ %
- [ whd in ⊢ (??%?); >atml_q0_q1 //
- | whd in ⊢ ((??%%)→?); #Hfalse destruct ]
- | #_ % // %2 @(ex_intro ?? c) % // ]
- ]
- | @(ex_intro ?? (mk_config ?? atm2 (mk_tape ? (tail ? ls) (option_hd ? ls) (c::rs))))
- %
- [ %
- [ whd in ⊢ (??%?); >atml_q0_q2 //
- | #_ @(ex_intro ?? ls) @(ex_intro ?? c) @(ex_intro ?? rs)
- % // % //
- ]
- | #Hfalse @False_ind @(absurd ?? Hfalse) %
- ]
- ]
-]
-qed.
-
-definition R_adv_to_mark_l ≝ λalpha,test,t1,t2.
- (current ? t1 = None ? → t1 = t2) ∧
- ∀ls,c,rs.
- (t1 = midtape alpha ls c rs →
- ((test c = true → t2 = t1) ∧
- (test c = false →
- (∀ls1,b,ls2. ls = ls1@b::ls2 →
- test b = true → (∀x.memb ? x ls1 = true → test x = false) →
- t2 = midtape ? ls2 b (reverse ? ls1@c::rs)) ∧
- ((∀x.memb ? x ls = true → test x = false) →
- ∀a,l. reverse ? (c::ls) = a::l → t2 = leftof ? a (l@rs))
- ))).
-
-definition adv_to_mark_l ≝
- λalpha,test.whileTM alpha (atml_step alpha test) atm2.
-
-lemma wsem_adv_to_mark_l :
- ∀alpha,test.
- WRealize alpha (adv_to_mark_l alpha test) (R_adv_to_mark_l alpha test).
-#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)
-[ * #Htapea *
- [ #H1 %
- [#_ @Htapea
- |#ls #c #rs #H2 >H2 in H1; whd in ⊢ (??%? → ?);
- #Hfalse destruct (Hfalse)
- ]
- | * #a * #Ha #Htest %
- [>Ha #H destruct (H);
- |#ls #c #rs #H2 %
- [#Hc <Htapea //
- |#Hc @False_ind >H2 in Ha; whd in ⊢ ((??%?)→?);
- #H destruct (H) /2/
- ]
- ]
- ]
-| #tapeb #tapec #Hleft #Hright #IH #HRfalse
- lapply (IH HRfalse) -IH #IH %
- [cases Hleft #ls0 * #a0 * #rs0 * * #Htapea #_ #_ >Htapea
- whd in ⊢ ((??%?)→?); #H destruct (H)
- |#ls #c #rs #Htapea %
- [#Hc cases Hleft #ls0 * #a0 * #rs0 * * #Htapea' #Htest @False_ind
- >Htapea' in Htapea; #H destruct /2/
- |cases Hleft #ls0 * #a * #rs0 *
- * #Htapea1 >Htapea in Htapea1; #H destruct (H) #_ #Htapeb
- #Hc %
- [*
- [#b #ls2 #Hls >Hls in Htapeb; #Htapeb #Htestb #_
- cases (proj2 ?? IH … Htapeb) #H1 #_ >H1 // >Htapeb %
- |#l1 #ls1 #b #ls2 #Hls >Hls in Htapeb; #Htapeb #Htestb #Hmemb
- cases (proj2 ?? IH … Htapeb) #_ #H1 >reverse_cons >associative_append
- @(proj1 ?? (H1 ?) … (refl …) Htestb …)
- [@Hmemb @memb_hd
- |#x #memx @Hmemb @memb_cons @memx
- ]
- ]
- |cases ls0 in Htapeb; normalize in ⊢ (%→?);
- [#Htapeb #Htest #a0 #l whd in ⊢ ((??%?)→?); #Hrev destruct (Hrev)
- >Htapeb in IH; #IH cases (proj1 ?? IH … (refl …)) //
- |#l1 #ls1 #Htapeb
- cases (proj2 ?? IH … Htapeb) #_ #H1 #Htest #a0 #l
- <(reverse_reverse … l) cases (reverse … l)
- [#H cut (a::l1::ls1 = [a0])
- [<(reverse_reverse … (a::l1::ls1)) >H //]
- #Hrev destruct (Hrev)
- |#a1 #l2 >reverse_cons >reverse_cons >reverse_cons
- #Hrev cut ([a] = [a1])
- [@(append_l2_injective_r ?? (a0::reverse … l2) … Hrev) //]
- #Ha <Ha >associative_append @(proj2 ?? (H1 ?))
- [@Htest @memb_hd
- |#x #membx @Htest @memb_cons @membx
- |<(append_l1_injective_r ?? (a0::reverse … l2) … Hrev) //
- ]
- ]
- ]
- ]
- ]
- ]
-qed.
-
-lemma terminate_adv_to_mark_l :
- ∀alpha,test.
- ∀t.Terminate alpha (adv_to_mark_l alpha test) t.
-#alpha #test #t
-@(terminate_while … (sem_atml_step alpha test))
- [ %
- | cases t
- [ % #t1 * #ls0 * #c0 * #rs0 * * #Hfalse destruct (Hfalse)
- |2,3: #a0 #al0 % #t1 * #ls0 * #c0 * #rs0 * * #Hfalse destruct (Hfalse)
- | #ls elim ls
- [#c #rs % #t1 * #ls0 * #c0 * #rs0 * *
- #H1 destruct (H1) #Hc0 #Ht1 normalize in Ht1;
- % #t2 * #ls1 * #c1 * #rs1 * * >Ht1
- normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
- | #rs0 #r0 #IH #ls #c % #t1 * #ls0 * #c0 * #rs0 * *
- #H1 destruct (H1) #Hc0 #Ht1 normalize in Ht1;
- >Ht1 @IH
- ]
- ]
- ]
-qed.
-
-lemma sem_adv_to_mark_l :
- ∀alpha,test.
- Realize alpha (adv_to_mark_l alpha test) (R_adv_to_mark_l alpha test).
-/2/
-qed.
-
-(*
- ADVANCE BOTH MARKS machine
-
- l1 does not contain marks ⇒
-
-
- input:
- l0 x* a l1 x0* a0 l2
- ^
-
- output:
- l0 x a* l1 x0 a0* l2
- ^
-*)
-
-definition adv_both_marks ≝ λalpha.
- adv_mark_r alpha · move_l ? ·
- adv_to_mark_l (FinProd alpha FinBool) (is_marked alpha) ·
- adv_mark_r alpha.
-
-definition R_adv_both_marks ≝ λalpha,t1,t2.
- ∀ls,x0,rs.
- t1 = midtape (FinProd … alpha FinBool) ls 〈x0,true〉 rs →
- (rs = [ ] → (* first case: rs empty *)
- t2 = rightof (FinProd … alpha FinBool) 〈x0,false〉 ls) ∧
- (∀l0,x,a,a0,b,l1,l1',l2.
- ls = (l1@〈x,true〉::l0) →
- (∀c.memb ? c l1 = true → is_marked ? c = false) →
- rs = (〈a0,b〉::l2) →
- reverse ? (〈x0,false〉::l1) = 〈a,false〉::l1' →
- t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (l1'@〈a0,true〉::l2)).
-
-lemma sem_adv_both_marks :
- ∀alpha.Realize ? (adv_both_marks alpha) (R_adv_both_marks alpha).
-#alpha
-@(sem_seq_app … (sem_adv_mark_r …)
- (sem_seq ????? (sem_move_l …)
- (sem_seq ????? (sem_adv_to_mark_l ? (is_marked ?))
- (sem_adv_mark_r alpha))) …)
-#intape #outtape * #tapea * #Hta * #tb * #Htb * #tc * #Htc #Hout
-#ls #c #rs #Hintape %
- [#Hrs >Hrs in Hintape; #Hintape lapply (proj2 ?? (proj1 ?? Hta … ) … Hintape) -Hta #Hta
- lapply (proj1 … Htb) >Hta -Htb #Htb lapply (Htb (refl …)) -Htb #Htb
- lapply (proj1 ?? Htc) <Htb -Htc #Htc lapply (Htc (refl …)) -Htc #Htc
- @sym_eq >Htc @(proj2 ?? Hout …) <Htc %
- |#l0 #x #a #a0 #b #l1 #l1' #l2 #Hls #Hmarks #Hrs #Hrev
- >Hrs in Hintape; >Hls #Hintape
- @(proj1 ?? (proj1 ?? Hout … ) ? false) -Hout
- lapply (proj1 … (proj1 … Hta …) … Hintape) #Htapea
- lapply (proj2 … Htb … Htapea) -Htb
- whd in match (mk_tape ????) ; #Htapeb
- lapply (proj1 ?? (proj2 ?? (proj2 ?? Htc … Htapeb) (refl …))) -Htc #Htc
- change with ((?::?)@?) in match (cons ???); <Hrev >reverse_cons
- >associative_append @Htc [%|%|@Hmarks]
- ]
-qed.
-
-(*
-definition R_adv_both_marks ≝
- λalpha,t1,t2.
- ∀l0,x,a,l1,x0,a0,l2. (∀c.memb ? c l1 = true → is_marked ? c = false) →
- (t1 = midtape (FinProd … alpha FinBool)
- (l1@〈a,false〉::〈x,true〉::l0) 〈x0,true〉 (〈a0,false〉::l2) →
- t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (reverse ? l1@〈x0,false〉::〈a0,true〉::l2)) ∧
- (t1 = midtape (FinProd … alpha FinBool)
- (l1@〈a,false〉::〈x,true〉::l0) 〈x0,true〉 [] →
- t2 = rightof ? 〈x0,false〉 (l1@〈a,false〉::〈x,true〉::l0)).
-
-lemma sem_adv_both_marks :
- ∀alpha.Realize ? (adv_both_marks alpha) (R_adv_both_marks alpha).
-#alpha #intape
-cases (sem_seq ????? (sem_adv_mark_r …)
- (sem_seq ????? (sem_move_l …)
- (sem_seq ????? (sem_adv_to_mark_l ? (is_marked ?))
- (sem_adv_mark_r alpha))) intape)
-#k * #outc * #Hloop whd in ⊢ (%→?);
-* #ta * whd in ⊢ (%→?); #Hs1 * #tb * whd in ⊢ (%→?); #Hs2
-* #tc * whd in ⊢ (%→%→?); #Hs3 #Hs4
-@(ex_intro ?? k) @(ex_intro ?? outc) %
-[ @Hloop
-| -Hloop #l0 #x #a #l1 #x0 #a0 #l2 #Hl1 #Hintape
- @(Hs4 … false) -Hs4
- lapply (Hs1 … Hintape) #Hta
- lapply (proj2 … Hs2 … Hta) #Htb
- cases (Hs3 … Htb)
- [ * #Hfalse normalize in Hfalse; destruct (Hfalse)
- | * #_ -Hs3 #Hs3
- lapply (Hs3 (l1@[〈a,false〉]) 〈x,true〉 l0 ???)
- [ #x1 #Hx1 cases (memb_append … Hx1)
- [ @Hl1
- | #Hx1' >(memb_single … Hx1') % ]
- | %
- | >associative_append %
- | >reverse_append #Htc @Htc ]
- ]
-qed. *)
-
-(*
- MATCH AND ADVANCE(f)
-
- input:
- l0 x* a l1 x0* a0 l2
- ^
-
- output:
- l0 x a* l1 x0 a0* l2 (f(x0) == true)
- ^
- l0 x* a l1 x0* a0 l2 (f(x0) == false)
- ^
-*)
-
-definition match_and_adv ≝
- λalpha,f.ifTM ? (test_char ? f)
- (adv_both_marks alpha) (clear_mark ?) tc_true.
-
-definition R_match_and_adv ≝
- λalpha,f,t1,t2.
- ∀ls,x0,rs.
- t1 = midtape (FinProd … alpha FinBool) ls 〈x0,true〉 rs →
- ((* first case: (f 〈x0,true〉 = false) *)
- (f 〈x0,true〉 = false) →
- t2 = midtape (FinProd … alpha FinBool) ls 〈x0,false〉 rs) ∧
- ((f 〈x0,true〉 = true) → rs = [ ] → (* second case: rs empty *)
- t2 = rightof (FinProd … alpha FinBool) 〈x0,false〉 ls) ∧
- ((f 〈x0,true〉 = true) →
- ∀l0,x,a,a0,b,l1,l1',l2.
- (* third case: we expect to have a mark on the left! *)
- ls = (l1@〈x,true〉::l0) →
- (∀c.memb ? c l1 = true → is_marked ? c = false) →
- rs = 〈a0,b〉::l2 →
- reverse ? (〈x0,false〉::l1) = 〈a,false〉::l1' →
- t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (l1'@〈a0,true〉::l2)).
-
-lemma sem_match_and_adv :
- ∀alpha,f.Realize ? (match_and_adv alpha f) (R_match_and_adv alpha f).
-#alpha #f #intape
-cases (sem_if ? (test_char ? f) … tc_true (sem_test_char ? f) (sem_adv_both_marks alpha) (sem_clear_mark ?) intape)
-#k * #outc * #Hloop #Hif @(ex_intro ?? k) @(ex_intro ?? outc)
-% [ @Hloop ] -Hloop
-(*
-@(sem_if_app … (sem_test_char ? f) (sem_adv_both_marks alpha) (sem_clear_mark ?))
-#intape #outape #Htb * #H *)
-cases Hif -Hif
-[ * #ta * whd in ⊢ (%→%→?); * * #c * #Hcurrent #fc #Hta #Houtc
- #ls #x #rs #Hintape >Hintape in Hcurrent; whd in ⊢ ((??%?)→?); #H destruct (H) %
- [%[>fc #H destruct (H)
- |#_ #Hrs >Hrs in Hintape; #Hintape >Hintape in Hta; #Hta
- cases (Houtc … Hta) #Houtc #_ @Houtc //
- ]
- |#l0 #x0 #a #a0 #b #l1 #l1' #l2 #Hls #Hmarks #Hrs #Hrev >Hintape in Hta; #Hta
- @(proj2 ?? (Houtc … Hta) … Hls Hmarks Hrs Hrev)
- ]
-| * #ta * * #H #Hta * #_ #Houtc #ls #c #rs #Hintape
- >Hintape in H; #H lapply(H … (refl …)) #fc %
- [%[#_ >Hintape in Hta; #Hta @(Houtc … Hta)
- |>fc #H destruct (H)
- ]
- |>fc #H destruct (H)
- ]
-]
-qed.
-
-definition R_match_and_adv_of ≝
- λalpha,t1,t2.current (FinProd … alpha FinBool) t1 = None ? → t2 = t1.
-
-lemma sem_match_and_adv_of :
- ∀alpha,f.Realize ? (match_and_adv alpha f) (R_match_and_adv_of alpha).
-#alpha #f #intape
-cases (sem_if ? (test_char ? f) … tc_true (sem_test_char ? f) (sem_adv_both_marks alpha) (sem_clear_mark ?) intape)
-#k * #outc * #Hloop #Hif @(ex_intro ?? k) @(ex_intro ?? outc)
-% [ @Hloop ] -Hloop
-cases Hif
-[ * #ta * whd in ⊢ (%→%→?); #Hta #Houtc #Hcur
- cases Hta * #x >Hcur * #Hfalse destruct (Hfalse)
-| * #ta * whd in ⊢ (%→%→?); * #_ #Hta * #Houtc #_ <Hta #Hcur >(Houtc Hcur) % ]
-qed.
-
-lemma sem_match_and_adv_full :
- ∀alpha,f.Realize ? (match_and_adv alpha f)
- (R_match_and_adv alpha f ∩ R_match_and_adv_of alpha).
-#alpha #f #intape cases (sem_match_and_adv ? f intape)
-#i * #outc * #Hloop #HR1 %{i} %{outc} % // % //
-cases (sem_match_and_adv_of ? f intape) #i0 * #outc0 * #Hloop0 #HR2
->(loop_eq … Hloop Hloop0) //
-qed.
-
-definition comp_step_subcase ≝ λalpha,c,elseM.
- ifTM ? (test_char ? (λx.x == c))
- (move_r … · adv_to_mark_r ? (is_marked alpha) · match_and_adv ? (λx.x == c))
- elseM tc_true.
-
-definition R_comp_step_subcase ≝
- λalpha,c,RelseM,t1,t2.
- ∀ls,x,rs.t1 = midtape (FinProd … alpha FinBool) ls 〈x,true〉 rs →
- (〈x,true〉 = c →
- ((* test true but no marks in rs *)
- (∀c.memb ? c rs = true → is_marked ? c = false) →
- ∀a,l. (a::l) = reverse ? (〈x,true〉::rs) →
- t2 = rightof (FinProd alpha FinBool) a (l@ls)) ∧
- ∀l1,x0,l2.
- (∀c.memb ? c l1 = true → is_marked ? c = false) →
- rs = l1@〈x0,true〉::l2 →
- (x = x0 →
- l2 = [ ] → (* test true but l2 is empty *)
- t2 = rightof ? 〈x0,false〉 ((reverse ? l1)@〈x,true〉::ls)) ∧
- (x = x0 →
- ∀a,a0,b,l1',l2'. (* test true and l2 is not empty *)
- 〈a,false〉::l1' = l1@[〈x0,false〉] →
- l2 = 〈a0,b〉::l2' →
- t2 = midtape ? (〈x,false〉::ls) 〈a,true〉 (l1'@〈a0,true〉::l2')) ∧
- (x ≠ x0 →(* test false *)
- t2 = midtape (FinProd … alpha FinBool) ((reverse ? l1)@〈x,true〉::ls) 〈x0,false〉 l2)) ∧
- (〈x,true〉 ≠ c → RelseM t1 t2).
-
-lemma dec_marked: ∀alpha,rs.
- decidable (∀c.memb ? c rs = true → is_marked alpha c = false).
-#alpha #rs elim rs
- [%1 #n normalize #H destruct
- |#a #tl cases (true_or_false (is_marked ? a)) #Ha
- [#_ %2 % #Hall @(absurd ?? not_eq_true_false) <Ha
- @Hall @memb_hd
- |* [#Hall %1 #c #memc cases (orb_true_l … memc)
- [#eqca >(\P eqca) @Ha |@Hall]
- |#Hnall %2 @(not_to_not … Hnall) #Hall #c #memc @Hall @memb_cons //
- ]
- qed.
-
-(* axiom daemon:∀P:Prop.P. *)
-
-lemma sem_comp_step_subcase :
- ∀alpha,c,elseM,RelseM.
- Realize ? elseM RelseM →
- Realize ? (comp_step_subcase alpha c elseM)
- (R_comp_step_subcase alpha c RelseM).
-#alpha #c #elseM #RelseM #Helse #intape
-cases (sem_if ? (test_char ? (λx.x == c)) … tc_true
- (sem_test_char ? (λx.x == c))
- (sem_seq ????? (sem_move_r …)
- (sem_seq ????? (sem_adv_to_mark_r ? (is_marked alpha))
- (sem_match_and_adv_full ? (λx.x == c)))) Helse intape)
-#k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc)
-% [ @Hloop ] -Hloop cases HR -HR
- [* #ta * whd in ⊢ (%→?); * * #cin * #Hcin #Hcintrue #Hta * #tb * whd in ⊢ (%→?); #Htb
- * #tc * whd in ⊢ (%→?); #Htc * whd in ⊢ (%→%→?); #Houtc #Houtc1
- #ls #x #rs #Hintape >Hintape in Hcin; whd in ⊢ ((??%?)→?); #H destruct (H) %
- [#_ cases (dec_marked ? rs) #Hdec
- [%
- [#_ #a #l1
- >Hintape in Hta; #Hta
- lapply (proj2 ?? Htb … Hta) -Htb -Hta cases rs in Hdec;
- (* by cases on rs *)
- [#_ whd in ⊢ ((???%)→?); #Htb >Htb in Htc; #Htc
- lapply (proj1 ?? Htc (refl …)) -Htc #Htc <Htc in Houtc1; #Houtc1
- normalize in ⊢ (???%→?); #Hl1 destruct(Hl1) @(Houtc1 (refl …))
- |#r0 #rs0 #Hdec whd in ⊢ ((???%)→?); #Htb >Htb in Htc; #Htc
- >reverse_cons >reverse_cons #Hl1
- cases (proj2 ?? Htc … (refl …))
- [* >(Hdec …) [ #Hfalse destruct(Hfalse) ] @memb_hd
- |* #_ -Htc #Htc cut (∃l2.l1 = l2@[〈x,true〉])
- [generalize in match Hl1; -Hl1 <(reverse_reverse … l1)
- cases (reverse ? l1)
- [#Hl1 cut ([a]=〈x,true〉::r0::rs0)
- [ <(reverse_reverse … (〈x,true〉::r0::rs0))
- >reverse_cons >reverse_cons <Hl1 %
- | #Hfalse destruct(Hfalse)]
- |#a0 #l10 >reverse_cons #Heq
- lapply (append_l2_injective_r ? (a::reverse ? l10) ???? Heq) //
- #Ha0 destruct(Ha0) /2/ ]
- |* #l2 #Hl2 >Hl2 in Hl1; #Hl1
- lapply (append_l1_injective_r ? (a::l2) … Hl1) // -Hl1 #Hl1
- >reverse_cons in Htc; #Htc lapply (Htc … (sym_eq … Hl1))
- [ #x0 #Hmemb @Hdec @memb_cons @Hmemb ]
- -Htc #Htc >Htc in Houtc1; #Houtc1 >associative_append @Houtc1 %
- ]
- ]
- ]
- |#l1 #x0 #l2 #_ #Hrs @False_ind
- @(absurd ?? not_eq_true_false)
- change with (is_marked ? 〈x0,true〉) in match true;
- @Hdec >Hrs @memb_append_l2 @memb_hd
- ]
- |% [#H @False_ind @(absurd …H Hdec)]
- (* by cases on l1 *) *
- [#x0 #l2 #Hdec normalize in ⊢ (%→?); #Hrs >Hrs in Hintape; #Hintape
- >Hintape in Hta; (* * * #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) #Hx *)
- #Hta lapply (proj2 … Htb … Hta) -Htb -Hta
- whd in match (mk_tape ????); whd in match (tail ??); #Htb cases Htc -Htc
- #_ #Htc cases (Htc … Htb) -Htc
- [2: * * #Hfalse normalize in Hfalse; destruct (Hfalse) ]
- * * #Htc >Htb in Htc; -Htb #Htc cases (Houtc … Htc) -Houtc *
- #H1 #H2 #H3 cases (true_or_false (x==x0)) #eqxx0
- [>(\P eqxx0) % [2: #H @False_ind /2/] %
- [#_ #Hl2 >(H2 … Hl2) <(\P eqxx0) [% | @Hcintrue]
- |#_ #a #a0 #b #l1' #l2' normalize in ⊢ (%→?); #Hdes destruct (Hdes)
- #Hl2 @(H3 … Hdec … Hl2) <(\P eqxx0) [@Hcintrue | % | @reverse_single]
- ]
- |% [% #eqx @False_ind lapply (\Pf eqxx0) #Habs @(absurd … eqx Habs)]
- #_ @H1 @(\bf ?) @(not_to_not ??? (\Pf eqxx0)) <(\P Hcintrue)
- #Hdes destruct (Hdes) %
- ]
- |#l1hd #l1tl #x0 #l2 #Hdec normalize in ⊢ (%→?); #Hrs >Hrs in Hintape; #Hintape
- >Hintape in Hta; (* * * #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) #Hx *)
- #Hta lapply (proj2 … Htb … Hta) -Htb -Hta
- whd in match (mk_tape ????); whd in match (tail ??); #Htb cases Htc -Htc
- #_ #Htc cases (Htc … Htb) -Htc
- [* #Hfalse @False_ind >(Hdec … (memb_hd …)) in Hfalse; #H destruct]
- * * #_ #Htc lapply (Htc … (refl …) (refl …) ?) -Htc
- [#x1 #membx1 @Hdec @memb_cons @membx1] #Htc
- cases (Houtc … Htc) -Houtc *
- #H1 #H2 #H3 #_ cases (true_or_false (x==x0)) #eqxx0
- [>(\P eqxx0) % [2: #H @False_ind /2/] %
- [#_ #Hl2 >(H2 … Hl2) <(\P eqxx0)
- [>reverse_cons >associative_append % | @Hcintrue]
- |#_ #a #a0 #b #l1' #l2' normalize in ⊢ (%→?); #Hdes (* destruct (Hdes) *)
- #Hl2 @(H3 ?????? (reverse … (l1hd::l1tl)) … Hl2) <(\P eqxx0)
- [@Hcintrue
- |>reverse_cons >associative_append %
- |#c0 #memc @Hdec <(reverse_reverse ? (l1hd::l1tl)) @memb_reverse @memc
- |>Hdes >reverse_cons >reverse_reverse >(\P eqxx0) %
- ]
- ]
- |% [% #eqx @False_ind lapply (\Pf eqxx0) #Habs @(absurd … eqx Habs)]
- #_ >reverse_cons >associative_append @H1 @(\bf ?)
- @(not_to_not ??? (\Pf eqxx0)) <(\P Hcintrue) #Hdes
- destruct (Hdes) %
- ]
- ]
- ]
- |>(\P Hcintrue) * #Hfalse @False_ind @Hfalse %
- ]
- | * #ta * * #Hcur #Hta #Houtc
- #l0 #x #rs #Hintape >Hintape in Hcur; #Hcur lapply (Hcur ? (refl …)) -Hcur #Hc %
- [ #Hfalse >Hfalse in Hc; #Hc cases (\Pf Hc) #Hc @False_ind @Hc %
- | -Hc #Hc <Hintape <Hta @Houtc ] ]
-qed.
-
-(*
-- se marcato, itero
-- se non è marcato
- + se è un bit, ho fallito il confronto della tupla corrente
- + se è un separatore, la tupla fa match
-
-
-ifTM ? (test_char ? is_marked)
- (single_finalTM … (comp_step_subcase unialpha 〈bit false,true〉
- (comp_step_subcase unialpha 〈bit true,true〉
- (clear_mark …))))
- (nop ?)
-*)
-
-definition comp_step ≝
- ifTM ? (test_char ? (is_marked ?))
- (single_finalTM … (comp_step_subcase FSUnialpha 〈bit false,true〉
- (comp_step_subcase FSUnialpha 〈bit true,true〉
- (comp_step_subcase FSUnialpha 〈null,true〉
- (clear_mark …)))))
- (nop ?)
- tc_true.
-
-(* da spostare *)
-
-lemma mem_append : ∀A,x,l1,l2. mem A x (l1@l2) →
- mem A x l1 ∨ mem A x l2.
-#A #x #l1 elim l1 normalize [/2/]
-#a #tl #Hind #l2 * [#eqxa %1 /2/ |#memx cases (Hind … memx) /3/]
-qed.
-
-let rec split_on A (l:list A) f acc on l ≝
- match l with
- [ nil ⇒ 〈acc,nil ?〉
- | cons a tl ⇒
- if f a then 〈acc,a::tl〉 else split_on A tl f (a::acc)
- ].
-
-lemma split_on_spec: ∀A:DeqSet.∀l,f,acc,res1,res2.
- split_on A l f acc = 〈res1,res2〉 →
- (∃l1. res1 = l1@acc ∧
- reverse ? l1@res2 = l ∧
- ∀x. memb ? x l1 =true → f x = false) ∧
- ∀a,tl. res2 = a::tl → f a = true.
-#A #l #f elim l
- [#acc #res1 #res2 normalize in ⊢ (%→?); #H destruct %
- [@(ex_intro … []) % normalize [% % | #x #H destruct]
- |#a #tl #H destruct
- ]
- |#a #tl #Hind #acc #res1 #res2 normalize in ⊢ (%→?);
- cases (true_or_false (f a)) #Hfa >Hfa normalize in ⊢ (%→?);
- #H destruct
- [% [@(ex_intro … []) % normalize [% % | #x #H destruct]
- |#a1 #tl1 #H destruct (H) //]
- |cases (Hind (a::acc) res1 res2 H) * #l1 * *
- #Hres1 #Htl #Hfalse #Htrue % [2:@Htrue] @(ex_intro … (l1@[a])) %
- [% [>associative_append @Hres1 | >reverse_append <Htl % ]
- |#x #Hmemx cases (memb_append ???? Hmemx)
- [@Hfalse | #H >(memb_single … H) //]
- ]
- ]
- ]
-qed.
-
-axiom mem_reverse: ∀A,l,x. mem A x (reverse ? l) → mem A x l.
-
-lemma split_on_spec_ex: ∀A:DeqSet.∀l,f.∃l1,l2.
- l1@l2 = l ∧ (∀x:A. memb ? x l1 = true → f x = false) ∧
- ∀a,tl. l2 = a::tl → f a = true.
-#A #l #f @(ex_intro … (reverse … (\fst (split_on A l f []))))
-@(ex_intro … (\snd (split_on A l f [])))
-cases (split_on_spec A l f [ ] ?? (eq_pair_fst_snd …)) * #l1 * *
->append_nil #Hl1 >Hl1 #Hl #Hfalse #Htrue %
- [% [@Hl|#x #memx @Hfalse <(reverse_reverse … l1) @memb_reverse //] | @Htrue]
-qed.
-
-(* versione esistenziale *)
-
-definition R_comp_step_true ≝ λt1,t2.
- ∃ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls 〈c,true〉 rs ∧
- ((* bit_or_null c = false *)
- (bit_or_null c = false → t2 = midtape ? ls 〈c,false〉 rs) ∧
- (* no marks in rs *)
- (bit_or_null c = true →
- (∀c.memb ? c rs = true → is_marked ? c = false) →
- ∀a,l. (a::l) = reverse ? (〈c,true〉::rs) →
- t2 = rightof (FinProd FSUnialpha FinBool) a (l@ls)) ∧
- (∀l1,c0,l2.
- bit_or_null c = true →
- (∀c.memb ? c l1 = true → is_marked ? c = false) →
- rs = l1@〈c0,true〉::l2 →
- (c = c0 →
- l2 = [ ] → (* test true but l2 is empty *)
- t2 = rightof ? 〈c0,false〉 ((reverse ? l1)@〈c,true〉::ls)) ∧
- (c = c0 →
- ∀a,a0,b,l1',l2'. (* test true and l2 is not empty *)
- 〈a,false〉::l1' = l1@[〈c0,false〉] →
- l2 = 〈a0,b〉::l2' →
- t2 = midtape ? (〈c,false〉::ls) 〈a,true〉 (l1'@〈a0,true〉::l2')) ∧
- (c ≠ c0 →(* test false *)
- t2 = midtape (FinProd … FSUnialpha FinBool)
- ((reverse ? l1)@〈c,true〉::ls) 〈c0,false〉 l2))).
-
-definition R_comp_step_false ≝
- λt1,t2.
- ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs →
- is_marked ? c = false ∧ t2 = t1.
-
-lemma is_marked_to_exists: ∀alpha,c. is_marked alpha c = true →
- ∃c'. c = 〈c',true〉.
-#alpha * #c * [#_ @(ex_intro … c) //| normalize #H destruct]
-qed.
-
-lemma exists_current: ∀alpha,c,t.
- current alpha t = Some alpha c → ∃ls,rs. t= midtape ? ls c rs.
-#alpha #c *
- [whd in ⊢ (??%?→?); #H destruct
- |#a #l whd in ⊢ (??%?→?); #H destruct
- |#a #l whd in ⊢ (??%?→?); #H destruct
- |#ls #c1 #rs whd in ⊢ (??%?→?); #H destruct
- @(ex_intro … ls) @(ex_intro … rs) //
- ]
-qed.
-
-lemma sem_comp_step :
- accRealize ? comp_step (inr … (inl … (inr … start_nop)))
- R_comp_step_true R_comp_step_false.
-@(acc_sem_if_app … (sem_test_char ? (is_marked ?))
- (sem_comp_step_subcase FSUnialpha 〈bit false,true〉 ??
- (sem_comp_step_subcase FSUnialpha 〈bit true,true〉 ??
- (sem_comp_step_subcase FSUnialpha 〈null,true〉 ??
- (sem_clear_mark …))))
- (sem_nop …) …)
-[#intape #outape #ta #Hta #Htb cases Hta * #cm * #Hcur
- cases (exists_current … Hcur) #ls * #rs #Hintape #cmark
- cases (is_marked_to_exists … cmark) #c #Hcm
- >Hintape >Hcm -Hintape -Hcm #Hta
- @(ex_intro … ls) @(ex_intro … c) @(ex_intro …rs) % [//] lapply Hta -Hta
- (* #ls #c #rs #Hintape whd in Hta;
- >Hintape in Hta; * #_ -Hintape forse non serve *)
- cases (true_or_false (c==bit false)) #Hc
- [>(\P Hc) #Hta %
- [%[whd in ⊢ ((??%?)→?); #Hdes destruct
- |#Hc @(proj1 ?? (proj1 ?? (Htb … Hta) (refl …)))
- ]
- |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (Htb … Hta) (refl …)))
- ]
- |cases (true_or_false (c==bit true)) #Hc1
- [>(\P Hc1) #Hta
- cut (〈bit true, true〉 ≠ 〈bit false, true〉) [% #Hdes destruct] #Hneq %
- [%[whd in ⊢ ((??%?)→?); #Hdes destruct
- |#Hc @(proj1 … (proj1 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) (refl …)))
- ]
- |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (proj2 ?? (Htb … Hta) Hneq … Hta)(refl …)))
- ]
- |cases (true_or_false (c==null)) #Hc2
- [>(\P Hc2) #Hta
- cut (〈null, true〉 ≠ 〈bit false, true〉) [% #Hdes destruct] #Hneq
- cut (〈null, true〉 ≠ 〈bit true, true〉) [% #Hdes destruct] #Hneq1 %
- [%[whd in ⊢ ((??%?)→?); #Hdes destruct
- |#Hc @(proj1 … (proj1 ?? (proj2 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) Hneq1 … Hta) (refl …)))
- ]
- |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (proj2 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) Hneq1 … Hta) (refl …)))
- ]
- |#Hta cut (bit_or_null c = false)
- [lapply Hc; lapply Hc1; lapply Hc2 -Hc -Hc1 -Hc2
- cases c normalize [* normalize /2/] /2/] #Hcut %
- [%[cases (Htb … Hta) #_ -Htb #Htb
- cases (Htb … Hta) [2: % #H destruct (H) normalize in Hc; destruct] #_ -Htb #Htb
- cases (Htb … Hta) [2: % #H destruct (H) normalize in Hc1; destruct] #_ -Htb #Htb
- lapply (Htb ?) [% #H destruct (H) normalize in Hc2; destruct]
- * #_ #Houttape #_ @(Houttape … Hta)
- |>Hcut #H destruct
- ]
- |#l1 #c0 #l2 >Hcut #H destruct
- ]
- ]
- ]
- ]
-|#intape #outape #ta #Hta #Htb #ls #c #rs #Hintape
- >Hintape in Hta; whd in ⊢ (%→?); * #Hmark #Hta % [@Hmark //]
- whd in Htb; >Htb //
-]
-qed.
-
-(* compare *)
-
-definition compare ≝
- whileTM ? comp_step (inr … (inl … (inr … start_nop))).
-
-(*
-definition R_compare :=
- λt1,t2.
- (t
-
- ∀ls,c,b,rs.t1 = midtape ? ls 〈c,b〉 rs →
- (b = true → rs = ....) →
- (b = false ∧ ....) ∨
- (b = true ∧
-
- rs = cs@l1@〈c0,true〉::cs0@l2
- (
-
-
- ls 〈c,b〉 cs l1 〈c0,b0〉 cs0 l2
-
-
-ACCETTAZIONE:
- ls (hd (Ls@〈grid,false〉))* (tail (Ls@〈grid,false〉)) l1 (hd (Ls@〈comma,false〉))* (tail (Ls@〈comma,false〉)) l2
- ^^^^^^^^^^^^^^^^^^^^^^^
-
- ls Ls 〈grid,false〉 l1 Ls 〈comma,true〉 l2
- ^^^^^^^^^^^^
-
-RIFIUTO: c ≠ d
-
- ls (hd (Ls@〈c,false〉))* (tail (Ls@〈c,false〉)) l1 (hd (Ls@〈d,false〉))* (tail (Ls@〈d,false〉)) l2
- ^^^^^^^^^^^^^^^^^^^^^^^
-
-
- ls Ls 〈c,true〉 l1 Ls 〈d,false〉 l2
- ^^^^^^^^
-
- ).
-
- |bs| = |b0s| →
- (∃la,d.〈b,true〉::bs = la@[〈grid,d〉] ∧ ∀x.memb ? x la → is_bit (\fst x) = true) →
- (∃lb,d0.〈b0,true〉::b0s = lb@[〈comma,d0〉] ∧ ∀x.memb ? x lb → is_bit (\fst x) = true) →
- t1 = midtape ? l0 〈b,true〉 (bs@l1@〈b0,true〉::b0s@l2 →
-
- mk_tape left (option current) right
-
- (b = grid ∧ b0 = comma ∧ bs = [] ∧ b0s = [] ∧
- t2 = midtape ? l0 〈grid,false〉 (l1@〈comma,true〉::l2)) ∨
- (b = bit x ∧ b = c ∧ bs = b0s
- *)
-
-definition list_cases2: ∀A.∀P:list A →list A →Prop.∀l1,l2. |l1| = |l2| →
-P [ ] [ ] → (∀a1,a2:A.∀tl1,tl2. |tl1| = |tl2| → P (a1::tl1) (a2::tl2)) → P l1 l2.
-#A #P #l1 #l2 #Hlen lapply Hlen @(list_ind2 … Hlen) //
-#tl1 #tl2 #hd1 #hd2 #Hind normalize #HlenS #H1 #H2 @H2 //
-qed.
-
-definition R_compare :=
- λt1,t2.
- ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs →
- (∀c'.bit_or_null c' = false → c = 〈c',true〉 → t2 = midtape ? ls 〈c',false〉 rs) ∧
- (∀c'. c = 〈c',false〉 → t2 = t1) ∧
-(* forse manca il caso no marks in rs *)
- ∀b,b0,bs,b0s,l1,l2.
- |bs| = |b0s| →
- (∀c.memb (FinProd … FSUnialpha FinBool) c bs = true → bit_or_null (\fst c) = true) →
- (∀c.memb (FinProd … FSUnialpha FinBool) c b0s = true → bit_or_null (\fst c) = true) →
- (∀c.memb ? c bs = true → is_marked ? c = false) →
- (∀c.memb ? c b0s = true → is_marked ? c = false) →
- (∀c.memb ? c l1 = true → is_marked ? c = false) →
- c = 〈b,true〉 → bit_or_null b = true →
- rs = bs@〈grid,false〉::l1@〈b0,true〉::b0s@〈comma,false〉::l2 →
- (〈b,true〉::bs = 〈b0,true〉::b0s ∧
- t2 = midtape ? (reverse ? bs@〈b,false〉::ls)
- 〈grid,false〉 (l1@〈b0,false〉::b0s@〈comma,true〉::l2)) ∨
- (∃la,c',d',lb,lc.c' ≠ d' ∧
- 〈b,false〉::bs = la@〈c',false〉::lb ∧
- 〈b0,false〉::b0s = la@〈d',false〉::lc ∧
- t2 = midtape (FinProd … FSUnialpha FinBool) (reverse ? la@
- reverse ? l1@
- 〈grid,false〉::
- reverse ? lb@
- 〈c',true〉::
- reverse ? la@ls)
- 〈d',false〉 (lc@〈comma,false〉::l2)).
-
-lemma wsem_compare : WRealize ? compare R_compare.
-#t #i #outc #Hloop
-lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%]
--Hloop * #t1 * #Hstar @(star_ind_l ??????? Hstar)
-[ whd in ⊢ (%→?); #Rfalse #ls #c #rs #Htapea %
- [ %
- [ #c' #Hc' #Hc lapply (Rfalse … Htapea) -Rfalse * >Hc
- whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse)
- | #c' #Hc lapply (Rfalse … Htapea) -Rfalse * #_
- #Htrue @Htrue ]
- | #b #b0 #bs #b0s #l1 #l2 #Hlen #Hbs1 #Hb0s1 #Hbs2 #Hb0s2 #Hl1 #Hc
- cases (Rfalse … Htapea) -Rfalse >Hc whd in ⊢ (??%?→?);#Hfalse destruct (Hfalse)
- ]
-| #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'
- [2: #Htapeb lapply (Htapeb Hc') -Htapeb #Htapeb #_ #_ %
- [%[#c1 #Hc1 #Heqc destruct (Heqc)
- cases (IH … Htapeb) * #_ #H #_ <Htapeb @(H … (refl…))
- |#c1 #Heqc destruct (Heqc)
- ]
- |#b #b0 #bs #b0s #l1 #l2 #_ #_ #_ #_ #_ #_
- #Heq destruct (Heq) >Hc' #Hfalse @False_ind destruct (Hfalse)
- ]
- |#_ (* no marks in rs ??? *) #_ #Hleft %
- [ %
- [ #c'' #Hc'' #Heq destruct (Heq) >Hc'' in Hc'; #H destruct (H)
- | #c0 #Hfalse destruct (Hfalse)
- ]
- |#b #b0 #bs #b0s #l1 #l2 #Hlen #Hbs1 #Hb0s1 #Hbs2 #Hb0s2 #Hl1
- #Heq destruct (Heq) #_ >append_cons; <associative_append #Hrs
- cases (Hleft … Hc' … Hrs) -Hleft
- [2: #c1 #memc1 cases (memb_append … memc1) -memc1 #memc1
- [cases (memb_append … memc1) -memc1 #memc1
- [@Hbs2 @memc1 |>(memb_single … memc1) %]
- |@Hl1 @memc1
- ]
- |* (* manca il caso in cui dopo una parte uguale il
- secondo nastro finisca ... ???? *)
- #_ cases (true_or_false (b==b0)) #eqbb0
- [2: #_ #Htapeb %2 lapply (Htapeb … (\Pf eqbb0)) -Htapeb #Htapeb
- cases (IH … Htapeb) * #_ #Hout #_
- @(ex_intro … []) @(ex_intro … b) @(ex_intro … b0)
- @(ex_intro … bs) @(ex_intro … b0s) %
- [%[%[@(\Pf eqbb0) | %] | %]
- |>(Hout … (refl …)) -Hout >Htapeb @eq_f3 [2,3:%]
- >reverse_append >reverse_append >associative_append
- >associative_append %
- ]
- |lapply Hbs1 lapply Hb0s1 lapply Hbs2 lapply Hb0s2 lapply Hrs
- -Hbs1 -Hb0s1 -Hbs2 -Hb0s2 -Hrs
- @(list_cases2 … Hlen)
- [#Hrs #_ #_ #_ #_ >associative_append >associative_append #Htapeb #_
- lapply (Htapeb … (\P eqbb0) … (refl …) (refl …)) -Htapeb #Htapeb
- cases (IH … Htapeb) -IH * #Hout #_ #_ %1 %
- [>(\P eqbb0) %
- |>(Hout grid (refl …) (refl …)) @eq_f
- normalize >associative_append %
- ]
- |* #a1 #ba1 * #a2 #ba2 #tl1 #tl2 #HlenS #Hrs #Hb0s2 #Hbs2 #Hb0s1 #Hbs1
- cut (ba1 = false) [@(Hbs2 〈a1,ba1〉) @memb_hd] #Hba1 >Hba1
- >associative_append >associative_append #Htapeb #_
- lapply (Htapeb … (\P eqbb0) … (refl …) (refl …)) -Htapeb #Htapeb
- cases (IH … Htapeb) -IH * #_ #_
- cut (ba2=false) [@(Hb0s2 〈a2,ba2〉) @memb_hd] #Hba2 >Hba2
- #IH cases(IH a1 a2 ?? (l1@[〈b0,false〉]) l2 HlenS ????? (refl …) ??)
- [3:#x #memx @Hbs1 @memb_cons @memx
- |4:#x #memx @Hb0s1 @memb_cons @memx
- |5:#x #memx @Hbs2 @memb_cons @memx
- |6:#x #memx @Hb0s2 @memb_cons @memx
- |7:#x #memx cases (memb_append …memx) -memx #memx
- [@Hl1 @memx | >(memb_single … memx) %]
- |8:@(Hbs1 〈a1,ba1〉) @memb_hd
- |9: >associative_append >associative_append %
- |-IH -Hbs1 -Hb0s1 -Hbs2 -Hrs *
- #Ha1a2 #Houtc %1 %
- [>(\P eqbb0) @eq_f destruct (Ha1a2) %
- |>Houtc @eq_f3
- [>reverse_cons >associative_append %
- |%
- |>associative_append %
- ]
- ]
- |-IH -Hbs1 -Hb0s1 -Hbs2 -Hrs *
- #la * #c' * #d' * #lb * #lc * * *
- #Hcd #H1 #H2 #Houtc %2
- @(ex_intro … (〈b,false〉::la)) @(ex_intro … c') @(ex_intro … d')
- @(ex_intro … lb) @(ex_intro … lc) %
- [%[%[@Hcd | >H1 %] |>(\P eqbb0) >Hba2 >H2 %]
- |>Houtc @eq_f3
- [>(\P eqbb0) >reverse_append >reverse_cons
- >reverse_cons >associative_append >associative_append
- >associative_append >associative_append %
- |%
- |%
- ]
- ]
- ]
- ]
- ]
- ]
- ]
- ]
-]
-qed.
-
-lemma WF_cst_niltape:
- WF ? (inv ? R_comp_step_true) (niltape (FinProd FSUnialpha FinBool)).
-@wf #t1 whd in ⊢ (%→?); * #ls * #c * #rs * #H destruct
-qed.
-
-lemma WF_cst_rightof:
- ∀a,ls. WF ? (inv ? R_comp_step_true) (rightof (FinProd FSUnialpha FinBool) a ls).
-#a #ls @wf #t1 whd in ⊢ (%→?); * #ls * #c * #rs * #H destruct
-qed.
-
-lemma WF_cst_leftof:
- ∀a,ls. WF ? (inv ? R_comp_step_true) (leftof (FinProd FSUnialpha FinBool) a ls).
-#a #ls @wf #t1 whd in ⊢ (%→?); * #ls * #c * #rs * #H destruct
-qed.
-
-lemma WF_cst_midtape_false:
- ∀ls,c,rs. WF ? (inv ? R_comp_step_true)
- (midtape (FinProd … FSUnialpha FinBool) ls 〈c,false〉 rs).
-#ls #c #rs @wf #t1 whd in ⊢ (%→?); * #ls' * #c' * #rs' * #H destruct
-qed.
-
-(* da spostare *)
-lemma not_nil_to_exists:∀A.∀l: list A. l ≠ [ ] →
- ∃a,tl. a::tl = l.
- #A * [* #H @False_ind @H // | #a #tl #_ @(ex_intro … a) @(ex_intro … tl) //]
- qed.
-
-lemma terminate_compare:
- ∀t. Terminate ? compare t.
-#t @(terminate_while … sem_comp_step) [%]
-cases t // #ls * #c * //
-#rs
-(* we cannot proceed by structural induction on the right tape,
- since compare moves the marks! *)
-cut (∃len. |rs| = len) [/2/]
-* #len lapply rs lapply c lapply ls -ls -c -rs elim len
- [#ls #c #rs #Hlen >(lenght_to_nil … Hlen) @wf #t1 whd in ⊢ (%→?); * #ls0 * #c0 * #rs0 * #Hmid destruct (Hmid)
- * * #H1 #H2 #_ cases (true_or_false (bit_or_null c0)) #Hc0
- [>(H2 Hc0 … (refl …)) // #x whd in ⊢ ((??%?)→?); #Hdes destruct
- |>(H1 Hc0) //
- ]
- |-len #len #Hind #ls #c #rs #Hlen @wf #t1 whd in ⊢ (%→?); * #ls0 * #c0 * #rs0 * #Hmid destruct (Hmid)
- * * #H1 #H2 #H3 cases (true_or_false (bit_or_null c0)) #Hc0
- [-H1 cases (split_on_spec_ex ? rs0 (is_marked ?)) #rs1 * #rs2
- cases rs2
- [(* no marks in right tape *)
- * * >append_nil #H >H -H #Hmarks #_
- cases (not_nil_to_exists ? (reverse (FSUnialpha×bool) (〈c0,true〉::rs0)) ?)
- [2: % >reverse_cons #H cases (nil_to_nil … H) #_ #H1 destruct]
- #a0 * #tl #H4 >(H2 Hc0 Hmarks a0 tl H4) //
- |(* the first marked element is a0 *)
- * #a0 #a0b #rs3 * * #H4 #H5 #H6 lapply (H3 ? a0 rs3 … Hc0 H5 ?)
- [<H4 @eq_f @eq_f2 [@eq_f @(H6 〈a0,a0b〉 … (refl …)) | %]
- |cases (true_or_false (c0==a0)) #eqc0a0 (* comparing a0 with c0 *)
- [* * (* we check if we have elements at the right of a0 *)
- lapply H4 -H4 cases rs3
- [#_ #Ht1 #_ #_ >(Ht1 (\P eqc0a0) (refl …)) //
- |(* a1 will be marked *)
- cases (not_nil_to_exists ? (rs1@[〈a0,false〉]) ?)
- [2: % #H cases (nil_to_nil … H) #_ #H1 destruct]
- * #a2 #a2b * #tl2 #H7 * #a1 #a1b #rs4 #H4 #_ #Ht1 #_
- cut (a2b =false)
- [lapply (memb_hd ? 〈a2,a2b〉 tl2) >H7 #mema2
- cases (memb_append … mema2)
- [@H5 |#H lapply(memb_single … H) #H2 destruct %]
- ]
- #Ha2b >Ha2b in H7; #H7
- >(Ht1 (\P eqc0a0) … H7 (refl …)) @Hind -Hind -Ht1 -Ha2b -H2 -H3 -H5 -H6
- <H4 in Hlen; >length_append normalize <plus_n_Sm #Hlen1
- >length_append normalize <(injective_S … Hlen1) @eq_f2 //
- cut (|〈a2,false〉::tl2|=|rs1@[〈a0,false〉]|) [>H7 %]
- >length_append normalize <plus_n_Sm <plus_n_O //
- ]
- |(* c0 =/= a0 *) * * #_ #_ #Ht1 >(Ht1 (\Pf eqc0a0)) //
- ]
- ]
- ]
- |>(H1 Hc0) //
- ]
-qed.
-
-lemma sem_compare : Realize ? compare R_compare.
-/2/ qed.
-
-(* new *)
-definition R_compare_new :=
- λt1,t2.
- ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs →
- (∀c'.bit_or_null c' = false → c = 〈c',true〉 → t2 = midtape ? ls 〈c',false〉 rs) ∧
- (∀c'. c = 〈c',false〉 → t2 = t1) ∧
-(* forse manca il caso no marks in rs *)
- ∀b,b0,bs,b0s,comma,l1,l2.
- |bs| = |b0s| →
- (∀c.memb (FinProd … FSUnialpha FinBool) c bs = true → bit_or_null (\fst c) = true) →
- (∀c.memb ? c bs = true → is_marked ? c = false) →
- (∀c.memb ? c b0s = true → is_marked ? c = false) →
- (∀c.memb ? c l1 = true → is_marked ? c = false) →
- c = 〈b,true〉 → bit_or_null b = true →
- rs = bs@〈grid,false〉::l1@〈b0,true〉::b0s@〈comma,false〉::l2 →
- (〈b,true〉::bs = 〈b0,true〉::b0s ∧
- t2 = midtape ? (reverse ? bs@〈b,false〉::ls)
- 〈grid,false〉 (l1@〈b0,false〉::b0s@〈comma,true〉::l2)) ∨
- (∃la,c',d',lb,lc.c' ≠ d' ∧
- 〈b,false〉::bs = la@〈c',false〉::lb ∧
- 〈b0,false〉::b0s = la@〈d',false〉::lc ∧
- t2 = midtape (FinProd … FSUnialpha FinBool) (reverse ? la@
- reverse ? l1@
- 〈grid,false〉::
- reverse ? lb@
- 〈c',true〉::
- reverse ? la@ls)
- 〈d',false〉 (lc@〈comma,false〉::l2)).
-
-lemma wsem_compare_new : WRealize ? compare R_compare_new.
-#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 %
- [ %
- [ #c' #Hc' #Hc lapply (Rfalse … Htapea) -Rfalse * >Hc
- whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse)
- | #c' #Hc lapply (Rfalse … Htapea) -Rfalse * #_
- #Htrue @Htrue ]
- | #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
- 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'
- [2: #Htapeb lapply (Htapeb Hc') -Htapeb #Htapeb #_ #_ %
- [%[#c1 #Hc1 #Heqc destruct (Heqc)
- cases (IH … Htapeb) * #_ #H #_ <Htapeb @(H … (refl…))
- |#c1 #Heqc destruct (Heqc)
- ]
- |#b #b0 #bs #b0s #comma #l1 #l2 #_ #_ #_ #_ #_
- #Heq destruct (Heq) >Hc' #Hfalse @False_ind destruct (Hfalse)
- ]
- |#_ (* no marks in rs ??? *) #_ #Hleft %
- [ %
- [ #c'' #Hc'' #Heq destruct (Heq) >Hc'' in Hc'; #H destruct (H)
- | #c0 #Hfalse destruct (Hfalse)
- ]
- |#b #b0 #bs #b0s #comma #l1 #l2 #Hlen #Hbs1 #Hbs2 #Hb0s2 #Hl1
- #Heq destruct (Heq) #_ >append_cons; <associative_append #Hrs
- cases (Hleft … Hc' … Hrs) -Hleft
- [2: #c1 #memc1 cases (memb_append … memc1) -memc1 #memc1
- [cases (memb_append … memc1) -memc1 #memc1
- [@Hbs2 @memc1 |>(memb_single … memc1) %]
- |@Hl1 @memc1
- ]
- |* (* manca il caso in cui dopo una parte uguale il
- secondo nastro finisca ... ???? *)
- #_ cases (true_or_false (b==b0)) #eqbb0
- [2: #_ #Htapeb %2 lapply (Htapeb … (\Pf eqbb0)) -Htapeb #Htapeb
- cases (IH … Htapeb) * #_ #Hout #_
- @(ex_intro … []) @(ex_intro … b) @(ex_intro … b0)
- @(ex_intro … bs) @(ex_intro … b0s) %
- [%[%[@(\Pf eqbb0) | %] | %]
- |>(Hout … (refl …)) -Hout >Htapeb @eq_f3 [2,3:%]
- >reverse_append >reverse_append >associative_append
- >associative_append %
- ]
- |lapply Hbs1 lapply Hbs2 lapply Hb0s2 lapply Hrs
- -Hbs1 -Hbs2 -Hb0s2 -Hrs
- @(list_cases2 … Hlen)
- [#Hrs #_ #_ #_ >associative_append >associative_append #Htapeb #_
- lapply (Htapeb … (\P eqbb0) … (refl …) (refl …)) -Htapeb #Htapeb
- cases (IH … Htapeb) -IH * #Hout #_ #_ %1 %
- [>(\P eqbb0) %
- |>(Hout grid (refl …) (refl …)) @eq_f
- normalize >associative_append %
- ]
- |* #a1 #ba1 * #a2 #ba2 #tl1 #tl2 #HlenS #Hrs #Hb0s2 #Hbs2 #Hbs1
- cut (ba1 = false) [@(Hbs2 〈a1,ba1〉) @memb_hd] #Hba1 >Hba1
- >associative_append >associative_append #Htapeb #_
- lapply (Htapeb … (\P eqbb0) … (refl …) (refl …)) -Htapeb #Htapeb
- cases (IH … Htapeb) -IH * #_ #_
- cut (ba2=false) [@(Hb0s2 〈a2,ba2〉) @memb_hd] #Hba2 >Hba2
- #IH cases(IH a1 a2 ??? (l1@[〈b0,false〉]) l2 HlenS ???? (refl …) ??)
- [4:#x #memx @Hbs1 @memb_cons @memx
- |5:#x #memx @Hbs2 @memb_cons @memx
- |6:#x #memx @Hb0s2 @memb_cons @memx
- |7:#x #memx cases (memb_append …memx) -memx #memx
- [@Hl1 @memx | >(memb_single … memx) %]
- |8:@(Hbs1 〈a1,ba1〉) @memb_hd
- |9: >associative_append >associative_append %
- |-IH -Hbs1 -Hbs2 -Hrs *
- #Ha1a2 #Houtc %1 %
- [>(\P eqbb0) @eq_f destruct (Ha1a2) %
- |>Houtc @eq_f3
- [>reverse_cons >associative_append %
- |%
- |>associative_append %
- ]
- ]
- |-IH -Hbs1 -Hbs2 -Hrs *
- #la * #c' * #d' * #lb * #lc * * *
- #Hcd #H1 #H2 #Houtc %2
- @(ex_intro … (〈b,false〉::la)) @(ex_intro … c') @(ex_intro … d')
- @(ex_intro … lb) @(ex_intro … lc) %
- [%[%[@Hcd | >H1 %] |>(\P eqbb0) >Hba2 >H2 %]
- |>Houtc @eq_f3
- [>(\P eqbb0) >reverse_append >reverse_cons
- >reverse_cons >associative_append >associative_append
- >associative_append >associative_append %
- |%
- |%
- ]
- ]
- ]
- ]
- ]
- ]
- ]
- ]
-]
-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/universal/tuples.ma".
-include "turing/universal/marks.ma".
-
-(*
-l0 x* a l1 x0* a0 l2 ------> l0 x a* l1 x0 a0* l2
- ^ ^
-
-if current (* x *) = #
- then
- else if x = 0
- then move_right; ----
- adv_to_mark_r;
- if current (* x0 *) = 0
- then advance_mark ----
- adv_to_mark_l;
- advance_mark
- else STOP
- else x = 1 (* analogo *)
-
-*)
-
-
-(*
- MARK NEXT TUPLE machine
- (partially axiomatized)
-
- marks the first character after the first bar (rightwards)
- *)
-
-definition mark_next_tuple ≝
- adv_to_mark_r ? bar_or_grid ·
- (ifTM ? (test_char ? (λc:STape.is_bar (\fst c)))
- (move_right_and_mark ?) (nop ?) tc_true).
-
-definition R_mark_next_tuple ≝
- λt1,t2.
- ∀ls,c,rs1,rs2.
- (* c non può essere un separatore ... speriamo *)
- t1 = midtape STape ls c (rs1@〈grid,false〉::rs2) →
- no_marks rs1 → no_grids rs1 → bar_or_grid c = false →
- (∃rs3,rs4,d,b.rs1 = rs3 @ 〈bar,false〉 :: rs4 ∧
- no_bars rs3 ∧
- Some ? 〈d,b〉 = option_hd ? (rs4@〈grid,false〉::rs2) ∧
- t2 = midtape STape (〈bar,false〉::reverse ? rs3@c::ls) 〈d,true〉 (tail ? (rs4@〈grid,false〉::rs2)))
- ∨
- (no_bars rs1 ∧ t2 = midtape ? (reverse ? rs1@c::ls) 〈grid,false〉 rs2).
-
-axiom daemon :∀P:Prop.P.
-
-axiom tech_split :
- ∀A:DeqSet.∀f,l.
- (∀x.memb A x l = true → f x = false) ∨
- (∃l1,c,l2.f c = true ∧ l = l1@c::l2 ∧ ∀x.memb ? x l1 = true → f x = false).
-(*#A #f #l elim l
-[ % #x normalize #Hfalse *)
-
-theorem sem_mark_next_tuple :
- Realize ? mark_next_tuple R_mark_next_tuple.
-#intape
-lapply (sem_seq ? (adv_to_mark_r ? bar_or_grid)
- (ifTM ? (test_char ? (λc:STape.is_bar (\fst c))) (move_right_and_mark ?) (nop ?) tc_true) ????)
-[@sem_if [5: // |6: @sem_move_right_and_mark |7: // |*:skip]
-| //
-|||#Hif cases (Hif intape) -Hif
- #j * #outc * #Hloop * #ta * #Hleft #Hright
- @(ex_intro ?? j) @ex_intro [|% [@Hloop] ]
- -Hloop
- #ls #c #rs1 #rs2 #Hrs #Hrs1 #Hrs1' #Hc
- cases (proj2 ?? Hleft … Hrs)
- [ * #Hfalse >Hfalse in Hc; #Htf destruct (Htf)
- | * * #_ #Hta #_ cases (tech_split STape (λc.is_bar (\fst c)) rs1)
- [ #H1 %2 % [@H1]
- lapply (Hta rs1 〈grid,false〉 rs2 (refl ??) ? ?)
- [ * #x #b #Hx whd in ⊢ (??%?); >(Hrs1' … Hx) >(H1 … Hx) %
- | %
- | -Hta #Hta cases Hright
- [ * #tb * whd in ⊢ (%→?); * * #c1 * >Hta
- whd in ⊢ ((??%?)→?); #H destruct (H) whd in ⊢ ((??%?)→?); #H destruct
- | * #tb * whd in ⊢ (%→?); * #_ #Htb >Htb >Hta
- whd in ⊢ (%→?); #H @H
- ]
- ]
- |* #rs3 * #c0 * #rs4 * * #Hc0 #Hsplit #Hrs3
- % @(ex_intro ?? rs3) @(ex_intro ?? rs4)
- lapply (Hta rs3 c0 (rs4@〈grid,false〉::rs2) ???)
- [#x #Hrs3' whd in ⊢ (??%?); >Hsplit in Hrs1;>Hsplit in Hrs3;
- #Hrs3 #Hrs1 >(Hrs1 …) [| @memb_append_l1 @Hrs3'|]
- >(Hrs3 … Hrs3') @Hrs1' >Hsplit @memb_append_l1 //
- |whd in ⊢ (??%?); >Hc0 %
- |>Hsplit >associative_append %
- ]-Hta #Hta
- cases Hright -Hright
- [* whd in ⊢ (%→?); #tb * * * #c1 * >Hta -Hta
- whd in ⊢ (??%?→?); #H destruct (H) #Hc1 #Htb
- whd in ⊢ (%→?); #Houtc
- cut (c1=〈bar,false〉)
- [lapply Hc1 lapply Hsplit cases c1 #c1l #c1r #Hsplit
- cases c1l normalize
- [#b #H destruct |2,3,5:#H destruct]
- #_ @eq_f @(Hrs1 … 〈c1l,c1r〉) >Hsplit @memb_append_l2 @memb_hd]
- #Hcut lapply Hsplit -Hsplit
- cases rs4 in Htb;
- [#Htb lapply(Houtc … Htb) -Houtc #Houtc #Hsplit
- @(ex_intro ?? grid) @(ex_intro ?? false) %
- [% [ % [<Hcut @Hsplit |@Hrs3 ] | % ]
- |>Houtc >Hcut %
- ]
- |* #r5l #r5r #rs5 #Htb
- lapply(Houtc … Htb) -Houtc #Houtc #Hsplit
- @(ex_intro ?? r5l) @(ex_intro ?? r5r) %
- [%[%[<Hcut @Hsplit| @Hrs3] | % ]
- |>Houtc >Hcut %
- ]
- ]
- |* whd in ⊢ (%→?); #tb * *
- #H @False_ind >Hta in H; #H lapply(H c0 (refl …))
- >Hc0 #H destruct
- ]
- ]
- ]
-]
-qed.
-
-definition init_current_on_match ≝
- move_l ? · adv_to_mark_l ? (λc:STape.is_grid (\fst c)) · move_r ? · mark ?.
-
-definition R_init_current_on_match ≝ λt1,t2.
- ∀l1,l2,c,rs. no_grids l1 → is_grid c = false →
- t1 = midtape STape (l1@〈c,false〉::〈grid,false〉::l2) 〈grid,false〉 rs →
- t2 = midtape STape (〈grid,false〉::l2) 〈c,true〉 ((reverse ? l1)@〈grid,false〉::rs).
-
-lemma sem_init_current_on_match :
- Realize ? init_current_on_match R_init_current_on_match.
-#intape
-cases (sem_seq ????? (sem_move_l ?)
- (sem_seq ????? (sem_adv_to_mark_l ? (λc:STape.is_grid (\fst c)))
- (sem_seq ????? (sem_move_r ?) (sem_mark ?))) intape)
-#k * #outc * #Hloop #HR
-@(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#l1 #l2 #c #rs #Hl1 #Hc #Hintape
-cases HR -HR #ta * whd in ⊢ (%→?); * #_ #Hta lapply (Hta … Hintape) -Hta -Hintape
-generalize in match Hl1; cases l1
- [#Hl1 whd in ⊢ ((???(??%%%))→?); #Hta
- * #tb * whd in ⊢ (%→?); * #_ #Htb cases (Htb … Hta) -Htb -Hta #_
- (* [* >Hc #Htemp destruct (Htemp) ] *)
- #Htb cases (Htb … Hc) -Htb #Htb #_
- lapply (Htb [ ] 〈grid,false〉 ? (refl ??) (refl …) Hl1)
- whd in ⊢ ((???(??%%%))→?); -Htb #Htb
- * #tc * whd in ⊢ (%→?); * #_ #Htc lapply (Htc … Htb) -Htb -Htc
- whd in ⊢ ((???(??%%%))→?); #Htc
- whd in ⊢ (%→?); * #Houtc #_ lapply (Houtc … Htc) -Houtc #Houtc
- >Houtc %
- |#d #tl #Htl whd in ⊢ ((???(??%%%))→?); #Hta
- * #tb * whd in ⊢ (%→?); * #_ #Htb cases (Htb … Hta) -Htb
- #_ (* [* >(Htl … (memb_hd …)) #Htemp destruct (Htemp)] *)
- #Htb cases (Htb ?) -Htb [2: @Htl @memb_hd]
- >append_cons #Htb #_ lapply (Htb … (refl ??) (refl …) ?)
- [#x #membx cases (memb_append … membx) -membx #membx
- [@Htl @memb_cons @membx | >(memb_single … membx) @Hc]]-Htb #Htb
- * #tc * whd in ⊢ (%→?); * #_ #Htc lapply (Htc … Htb) -Htb -Htc
- >reverse_append >associative_append whd in ⊢ ((???(??%%%))→?); #Htc
- whd in ⊢ (%→?); * #Houtc #_ lapply (Houtc … Htc) -Houtc #Houtc
- >Houtc >reverse_cons >associative_append %
- ]
-qed.
-
-(*
-definition init_current_gen ≝
- seq ? (adv_to_mark_l ? (is_marked ?))
- (seq ? (clear_mark ?)
- (seq ? (move_l ?)
- (seq ? (adv_to_mark_l ? (λc:STape.is_grid (\fst c)))
- (seq ? (move_r ?) (mark ?))))).
-
-definition R_init_current_gen ≝ λt1,t2.
- ∀l1,c,l2,b,l3,c1,rs,c0,b0. no_marks l1 → no_grids l2 →
- Some ? 〈c0,b0〉 = option_hd ? (reverse ? (〈c,true〉::l2)) →
- t1 = midtape STape (l1@〈c,true〉::l2@〈grid,b〉::l3) 〈c1,false〉 rs →
- t2 = midtape STape (〈grid,b〉::l3) 〈c0,true〉
- ((tail ? (reverse ? (l1@〈c,false〉::l2))@〈c1,false〉::rs)).
-
-lemma sem_init_current_gen : Realize ? init_current_gen R_init_current_gen.
-#intape
-cases (sem_seq ????? (sem_adv_to_mark_l ? (is_marked ?))
- (sem_seq ????? (sem_clear_mark ?)
- (sem_seq ????? (sem_move_l ?)
- (sem_seq ????? (sem_adv_to_mark_l ? (λc:STape.is_grid (\fst c)))
- (sem_seq ????? (sem_move_r ?) (sem_mark ?))))) intape)
-#k * #outc * #Hloop #HR
-@(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#l1 #c #l2 #b #l3 #c1 #rs #c0 #b0 #Hl1 #Hl2 #Hc #Hintape
-cases HR -HR #ta * whd in ⊢ (%→?); #Hta cases (Hta … Hintape) -Hta -Hintape
- [ * #Hfalse normalize in Hfalse; destruct (Hfalse) ]
-* #_ #Hta lapply (Hta l1 〈c,true〉 ? (refl ??) ??) [@Hl1|%] -Hta #Hta
-* #tb * whd in ⊢ (%→?); #Htb lapply (Htb … Hta) -Htb -Hta #Htb
-* #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htc -Htb
-generalize in match Hc; generalize in match Hl2; cases l2
- [#_ whd in ⊢ ((???%)→?); #Htemp destruct (Htemp)
- whd in ⊢ ((???(??%%%))→?); #Htc
- * #td * whd in ⊢ (%→?); #Htd cases (Htd … Htc) -Htd
- [2: * whd in ⊢ (??%?→?); #Htemp destruct (Htemp) ]
- * #_ #Htd >Htd in Htc; -Htd #Htd
- * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd) -Htd
- >reverse_append >reverse_cons
- whd in ⊢ ((???(??%%%))→?); #Hte
- whd in ⊢ (%→?); #Houtc lapply (Houtc … Hte) -Houtc -Hte #Houtc
- >Houtc %
- |#d #tl #Htl #Hc0 whd in ⊢ ((???(??%%%))→?); #Htc
- * #td * whd in ⊢ (%→?); #Htd cases (Htd … Htc) -Htd
- [* >(Htl … (memb_hd …)) whd in ⊢ (??%?→?); #Htemp destruct (Htemp)]
- * #Hd #Htd lapply (Htd … (refl ??) (refl ??) ?)
- [#x #membx @Htl @memb_cons @membx] -Htd #Htd
- * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd) -Htd
- >reverse_append >reverse_cons >reverse_cons
- >reverse_cons in Hc0; >reverse_cons cases (reverse ? tl)
- [normalize in ⊢ (%→?); #Hc0 destruct (Hc0) #Hte
- whd in ⊢ (%→?); #Houtc lapply (Houtc … Hte) -Houtc -Hte #Houtc
- >Houtc %
- |* #c2 #b2 #tl2 normalize in ⊢ (%→?); #Hc0 destruct (Hc0)
- whd in ⊢ ((???(??%%%))→?); #Hte
- whd in ⊢ (%→?); #Houtc lapply (Houtc … Hte) -Houtc -Hte #Houtc
- >Houtc >associative_append >associative_append >associative_append %
- ]
- ]
-qed.
-*)
-
-definition init_current ≝
- adv_to_mark_l ? (is_marked ?) ·clear_mark ? ·
- adv_to_mark_l ? (λc:STape.is_grid (\fst c)) · move_r ? · mark ?.
-
-definition R_init_current ≝ λt1,t2.
- ∀l1,c,l2,b,l3,c1,rs,c0,b0. no_marks l1 → no_grids l2 → is_grid c = false →
- Some ? 〈c0,b0〉 = option_hd ? (reverse ? (〈c,true〉::l2)) →
- t1 = midtape STape (l1@〈c,true〉::l2@〈grid,b〉::l3) 〈c1,false〉 rs →
- t2 = midtape STape (〈grid,b〉::l3) 〈c0,true〉
- ((tail ? (reverse ? (l1@〈c,false〉::l2))@〈c1,false〉::rs)).
-
-lemma sem_init_current : Realize ? init_current R_init_current.
-#intape
-cases (sem_seq ????? (sem_adv_to_mark_l ? (is_marked ?))
- (sem_seq ????? (sem_clear_mark ?)
- (sem_seq ????? (sem_adv_to_mark_l ? (λc:STape.is_grid (\fst c)))
- (sem_seq ????? (sem_move_r ?) (sem_mark ?)))) intape)
-#k * #outc * #Hloop #HR
-@(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop]
-cases HR -HR #ta * whd in ⊢ (%→?); * #_ #Hta
-* #tb * whd in ⊢ (%→?); * #_ #Htb
-* #tc * whd in ⊢ (%→?); * #_ #Htc
-* #td * whd in ⊢ (%→%→?); * #_ #Htd * #Houtc #_
-#l1 #c #l2 #b #l3 #c1 #rs #c0 #b0 #Hl1 #Hl2 #Hc #Hc0 #Hintape
-cases (Hta … Hintape) #_ -Hta #Hta cases (Hta (refl …)) -Hta
-#Hta #_ lapply (Hta l1 〈c,true〉 ? (refl ??) ??) [@Hl1|%]
--Hta #Hta lapply (Htb … Hta) -Htb #Htb cases (Htc … Htb) #_ #Htc
-cases (Htc Hc) -Htc #Htc #_ lapply (Htc … (refl ??) (refl ??) ?) [@Hl2]
--Htc #Htc lapply (Htd … Htc) -Htd
->reverse_append >reverse_cons
->reverse_cons in Hc0; cases (reverse … l2)
-[ normalize in ⊢ (%→?); #Hc0 destruct (Hc0)
- #Htd >(Houtc … Htd) %
-| * #c2 #b2 #tl2 normalize in ⊢ (%→?);
- #Hc0 #Htd >(Houtc … Htd)
- whd in ⊢ (???%); destruct (Hc0)
- >associative_append >associative_append %
-]
-qed.
-
-definition match_tuple_step ≝
- ifTM ? (test_char ? (λc:STape.¬ is_grid (\fst c)))
- (single_finalTM ?
- (compare ·
- (ifTM ? (test_char ? (λc:STape.is_grid (\fst c)))
- (nop ?)
- (mark_next_tuple ·
- (ifTM ? (test_char ? (λc:STape.is_grid (\fst c)))
- (mark ?) (move_l ? · init_current) tc_true)) tc_true)))
- (nop ?) tc_true.
-
-definition R_match_tuple_step_true_new ≝ λt1,t2.
- ∃ls,cur,rs.t1 = midtape STape ls cur rs \wedge
- \fst cur ≠ grid ∧
- (∀ls0,c,l1,l2,c1,l3,l4,rs0,n.
- only_bits_or_nulls l1 → no_marks l1 (* → no_grids l2 *) →
- bit_or_null c = true → bit_or_null c1 = true →
- only_bits_or_nulls l3 → S n = |l1| → |l1| = |l3| →
- table_TM (S n) (l2@〈c1,false〉::l3@〈comma,false〉::l4) →
- ls = 〈grid,false〉::ls0 → cur = 〈c,true〉 →
- rs = l1@〈grid,false〉::l2@〈c1,true〉::l3@〈comma,false〉::l4@〈grid,false〉::rs0 →
- (* facciamo match *)
- (〈c,false〉::l1 = 〈c1,false〉::l3 ∧
- t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls0) 〈grid,false〉
- (l2@〈c1,false〉::l3@〈comma,true〉::l4@〈grid,false〉::rs0))
- ∨
- (* non facciamo match e marchiamo la prossima tupla *)
- (〈c,false〉::l1 ≠ 〈c1,false〉::l3 ∧
- ∃c2,l5,l6.l4 = l5@〈bar,false〉::〈c2,false〉::l6 ∧
- (* condizioni su l5 l6 l7 *)
- t2 = midtape STape (〈grid,false〉::ls0) 〈c,true〉
- (l1@〈grid,false〉::l2@〈c1,false〉::l3@〈comma,false〉::
- l5@〈bar,false〉::〈c2,true〉::l6@〈grid,false〉::rs0))
- ∨
- (* non facciamo match e non c'è una prossima tupla:
- non specifichiamo condizioni sul nastro di output, perché
- non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
- (〈c,false〉::l1 ≠ 〈c1,false〉::l3 ∧ no_bars l4 ∧ current ? t2 = Some ? 〈grid,true〉)).
-
-(* universal version
-definition R_match_tuple_step_true ≝ λt1,t2.
- ∀ls,cur,rs.t1 = midtape STape ls cur rs →
- \fst cur ≠ grid ∧
- (∀ls0,c,l1,l2,c1,l3,l4,rs0,n.
- only_bits_or_nulls l1 → no_marks l1 (* → no_grids l2 *) →
- bit_or_null c = true → bit_or_null c1 = true →
- only_bits_or_nulls l3 → S n = |l1| → |l1| = |l3| →
- table_TM (S n) (l2@〈c1,false〉::l3@〈comma,false〉::l4) →
- ls = 〈grid,false〉::ls0 → cur = 〈c,true〉 →
- rs = l1@〈grid,false〉::l2@〈c1,true〉::l3@〈comma,false〉::l4@〈grid,false〉::rs0 →
- (* facciamo match *)
- (〈c,false〉::l1 = 〈c1,false〉::l3 ∧
- t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls0) 〈grid,false〉
- (l2@〈c1,false〉::l3@〈comma,true〉::l4@〈grid,false〉::rs0))
- ∨
- (* non facciamo match e marchiamo la prossima tupla *)
- (〈c,false〉::l1 ≠ 〈c1,false〉::l3 ∧
- ∃c2,l5,l6.l4 = l5@〈bar,false〉::〈c2,false〉::l6 ∧
- (* condizioni su l5 l6 l7 *)
- t2 = midtape STape (〈grid,false〉::ls0) 〈c,true〉
- (l1@〈grid,false〉::l2@〈c1,false〉::l3@〈comma,false〉::
- l5@〈bar,false〉::〈c2,true〉::l6@〈grid,false〉::rs0))
- ∨
- (* non facciamo match e non c'è una prossima tupla:
- non specifichiamo condizioni sul nastro di output, perché
- non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
- (〈c,false〉::l1 ≠ 〈c1,false〉::l3 ∧ no_bars l4 ∧ current ? t2 = Some ? 〈grid,true〉)).
-*)
-
-definition R_match_tuple_step_true ≝ λt1,t2.
- ∃ls,cur,rs.t1 = midtape STape ls cur rs \wedge
- \fst cur ≠ grid ∧
- (∀ls0,c,l1,l2,c1,l3,l4,rs0,n.
- only_bits_or_nulls l1 → no_marks l1 (* → no_grids l2 *) →
- bit_or_null c = true → bit_or_null c1 = true →
- only_bits_or_nulls l3 → S n = |l1| → |l1| = |l3| →
- table_TM (S n) (l2@〈c1,false〉::l3@〈comma,false〉::l4) →
- ls = 〈grid,false〉::ls0 → cur = 〈c,true〉 →
- rs = l1@〈grid,false〉::l2@〈c1,true〉::l3@〈comma,false〉::l4@〈grid,false〉::rs0 →
- (* facciamo match *)
- (〈c,false〉::l1 = 〈c1,false〉::l3 ∧
- t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls0) 〈grid,false〉
- (l2@〈c1,false〉::l3@〈comma,true〉::l4@〈grid,false〉::rs0))
- ∨
- (* non facciamo match e marchiamo la prossima tupla *)
- (〈c,false〉::l1 ≠ 〈c1,false〉::l3 ∧
- ∃c2,l5,l6.l4 = l5@〈bar,false〉::〈c2,false〉::l6 ∧
- (* condizioni su l5 l6 l7 *)
- t2 = midtape STape (〈grid,false〉::ls0) 〈c,true〉
- (l1@〈grid,false〉::l2@〈c1,false〉::l3@〈comma,false〉::
- l5@〈bar,false〉::〈c2,true〉::l6@〈grid,false〉::rs0))
- ∨
- (* non facciamo match e non c'è una prossima tupla:
- non specifichiamo condizioni sul nastro di output, perché
- non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
- (〈c,false〉::l1 ≠ 〈c1,false〉::l3 ∧ no_bars l4 ∧ current ? t2 = Some ? 〈grid,true〉)).
-
-definition R_match_tuple_step_false ≝ λt1,t2.
- ∀ls,c,rs.t1 = midtape STape ls c rs → is_grid (\fst c) = true ∧ t2 = t1.
-
-include alias "basics/logic.ma".
-
-(*
-lemma eq_f4: ∀A1,A2,A3,A4,B.∀f:A1 → A2 →A3 →A4 →B.
- ∀x1,x2,x3,x4,y1,y2,y3,y4. x1 = y1 → x2 = y2 →x3=y3 →x4 = y4 →
- f x1 x2 x3 x4 = f y1 y2 y3 y4.
-//
-qed-. *)
-
-lemma some_option_hd: ∀A.∀l:list A.∀a.∃b.
- Some ? b = option_hd ? (l@[a]) .
-#A #l #a cases l normalize /2/
-qed.
-
-axiom tech_split2 : ∀A,l1,l2,l3,l4,x.
- memb A x l1 = false → memb ? x l3 = false →
- l1@x::l2 = l3@x::l4 → l1 = l3 ∧ l2 = l4.
-
-axiom injective_append : ∀A,l.injective … (λx.append A x l).
-
-lemma sem_match_tuple_step:
- accRealize ? match_tuple_step (inr … (inl … (inr … start_nop)))
- R_match_tuple_step_true R_match_tuple_step_false.
-@(acc_sem_if_app … (sem_test_char ? (λc:STape.¬ is_grid (\fst c))) …
- (sem_seq … sem_compare
- (sem_if … (sem_test_char ? (λc:STape.is_grid (\fst c)))
- (sem_nop …)
- (sem_seq … sem_mark_next_tuple
- (sem_if … (sem_test_char ? (λc:STape.is_grid (\fst c)))
- (sem_mark ?) (sem_seq … (sem_move_l …) (sem_init_current …))))))
- (sem_nop ?) …)
-[(* is_grid: termination case *)
- 2:#t1 #t2 #t3 whd in ⊢ (%→?); * #Hc #H #H1 whd #ls #c #rs #Ht1 %
- [lapply(Hc c ?) [>Ht1 %] #Hgrid @injective_notb @Hgrid |>H1 @H]
-|#tapea #tapeout #tapeb whd in ⊢ (%→?); #Hcur
- * #tapec * whd in ⊢ (%→?); #Hcompare #Hor
- cases Hcur * #c * -Hcur #Hcur #Hgrid #Htapeb cases (current_to_midtape … Hcur)
- #ls * #rs #Htapea @(ex_intro … ls) @(ex_intro … c) @(ex_intro … rs) %
- [%[@Htapea | cases (true_or_false (\fst c == grid))
- [#eqc @False_ind >(\P eqc) in Hgrid; normalize #H destruct |#eqc @(\Pf eqc)]]]
- #ls0 #cur #l1 #l2 #c1 #l3 #l4 #rs0 #n #Hl1bitnull #Hl1marks #Hc #Hc1 #Hl3 #eqn
- #eqlen #Htable #Hls -Hcur #Hcur #Hrs >Htapea in Htapeb; >Hls >Hcur >Hrs #Htapeb
- cases (Hcompare … Htapeb) -Hcompare -Htapeb * #_ #_ #Hcompare
- cases (Hcompare cur c1 l1 l3 l2 (l4@〈grid,false〉::rs0) eqlen Hl1bitnull Hl3 Hl1marks … (refl …) Hc ?)
- -Hcompare
- [* #Htemp destruct (Htemp) #Htapec %1 % % [%]
- >Htapec in Hor; -Htapec *
- [2: * #t3 * whd in ⊢ (%→?); * #H #_ @False_ind
- lapply (H … (refl …)) whd in ⊢ ((??%?)→?); #H destruct (H)
- |* #taped * whd in ⊢ (%→?); * #_
- #Htaped whd in ⊢ (%→?); #Htapeout >Htapeout >Htaped
- %
- ]
- |* #la * #c' * #d' * #lb * #lc * * * #H1 #H2 #H3 #Htapec
- cut (〈cur,false〉::l1 ≠ 〈c1,false〉::l3)
- [>H2 >H3 elim la
- [@(not_to_not …H1) normalize #H destruct (H) %
- |#x #tl @not_to_not normalize #H destruct //
- ]
- ] #Hnoteq
- cut (bit_or_null d' = true)
- [cases la in H3;
- [normalize in ⊢ (%→?); #H destruct //
- |#x #tl #H @(Hl3 〈d',false〉)
- normalize in H; destruct @memb_append_l2 @memb_hd
- ]
- ] #Hd'
- >Htapec in Hor; -Htapec *
- [* #taped * whd in ⊢ (%→?); * * #c0 * normalize in ⊢ (%→?);
- #Hdes destruct (Hdes) >(bit_or_null_not_grid ? Hd') #Htemp destruct (Htemp)
- |* #taped * whd in ⊢ (%→?); * #_ (* * #_ #H cases (H … (refl …)) -H #_ *)
- #Htaped * #tapee * whd in ⊢ (%→?); #Htapee
- <(associative_append ? lc (〈comma,false〉::l4)) in Htaped; #Htaped
- cases (Htapee … Htaped ???) -Htaped -Htapee
- [* #rs3 * * (* we proceed by cases on rs4 *)
- [(* rs4 is empty : the case is absurd since the tape
- cannot end with a bar *)
- * #d * #b * * * #Heq1 @False_ind
- cut (∀A,l1,l2.∀a:A. a::l1@l2=(a::l1)@l2) [//] #Hcut
- >Hcut in Htable; >H3 >associative_append
- normalize >Heq1 <associative_append >Hcut
- <associative_append #Htable @(absurd … Htable)
- @last_of_table
- |(* rs4 not empty *)
- * #d2 #b2 #rs3' * #d * #b * * * #Heq1 #Hnobars
- cut (memb STape 〈d2,b2〉 (l2@〈c1,false〉::l3@〈comma,false〉::l4) = true)
- [@memb_append_l2
- cut (∀A,l1,l2.∀a:A. a::l1@l2=(a::l1)@l2) [//] #Hcut
- >Hcut >H3 >associative_append @memb_append_l2
- @memb_cons >Heq1 @memb_append_l2 @memb_cons @memb_hd] #d2intable
- cut (is_grid d2 = false)
- [@(no_grids_in_table … Htable … 〈d2,b2〉 d2intable)] #Hd2
- cut (b2 = false)
- [@(no_marks_in_table … Htable … 〈d2,b2〉 d2intable)] #Hb2
- >Hb2 in Heq1; #Heq1 -Hb2 -b2
- whd in ⊢ ((???%)→?); #Htemp destruct (Htemp) #Htapee >Htapee -Htapee *
- [(* we know current is not grid *)
- * #tapef * whd in ⊢ (%→?); * * #c0 *
- normalize in ⊢ (%→?); #Hdes destruct (Hdes) >Hd2
- #Htemp destruct (Htemp)
- |* #tapef * whd in ⊢ (%→?); * #_ #Htapef
- * #tapeg >Htapef -Htapef *
- (* move_l *)
- whd in ⊢ (%→?); * #_ #H lapply (H … (refl …)) whd in ⊢ (???%→?); -H #Htapeg
- >Htapeg -Htapeg
- (* init_current *)
- whd in ⊢ (%→?); #Htapeout
- cases (some_option_hd ? (reverse ? (reverse ? la)) 〈c',true〉)
- * #c00 #b00 #Hoption
- lapply
- (Htapeout (reverse ? rs3 @〈d',false〉::reverse ? la@reverse ? l2@(〈grid,false〉::reverse ? lb))
- c' (reverse ? la) false ls0 bar (〈d2,true〉::rs3'@〈grid,false〉::rs0) c00 b00 ?????) -Htapeout
- [whd in ⊢ (??(??%??)?); @eq_f3 [2:%|3: %]
- >associative_append
- generalize in match (〈c',true〉::reverse ? la@〈grid,false〉::ls0); #l
- whd in ⊢ (???(???%)); >associative_append >associative_append %
- |>reverse_cons @Hoption
- |cases la in H2;
- [normalize in ⊢ (%→?); #Htemp destruct (Htemp)
- @bit_or_null_not_grid @Hc
- |#x #tl normalize in ⊢ (%→?); #Htemp destruct (Htemp)
- @bit_or_null_not_grid @(Hl1bitnull 〈c',false〉) @memb_append_l2 @memb_hd
- ]
- |cut (only_bits_or_nulls (la@(〈c',false〉::lb)))
- [<H2 whd #c0 #Hmemb cases (orb_true_l … Hmemb)
- [#eqc0 >(\P eqc0) @Hc |@Hl1bitnull]
- |#Hl1' #x #Hx @bit_or_null_not_grid @Hl1'
- @memb_append_l1 @daemon
- ]
- |@daemon] #Htapeout % %2 % //
- @(ex_intro … d2)
- cut (∃rs32.rs3 = lc@〈comma,false〉::rs32)
- [ (*cases (tech_split STape (λc.c == 〈bar,false〉) l4)
- [
- | * #l41 * * #cbar #bfalse * #l42 * * #Hbar #Hl4 #Hl41
- @(ex_intro ?? l41) >Hl4 in Heq1; #Heq1
-
- cut (sublist … lc l3)
- [ #x #Hx cases la in H3;
- [ normalize #H3 destruct (H3) @Hx
- | #p #la' normalize #Hla' destruct (Hla')
- @memb_append_l2 @memb_cons @Hx ] ] #Hsublist*)
- @daemon]
- * #rs32 #Hrs3
- (* cut
- (〈c1,false〉::l3@〈comma,false〉::l4= la@〈d',false〉::rs3@〈bar,false〉::〈d2,b2〉::rs3')
- [@daemon] #Hcut *)
- cut (l4=rs32@〈bar,false〉::〈d2,false〉::rs3')
- [ >Hrs3 in Heq1; @daemon ] #Hl4
- @(ex_intro … rs32) @(ex_intro … rs3') % [@Hl4]
- >Htapeout @eq_f2
- [(* by Hoption, H2 *) @daemon
- |(*>Hrs3 *)>append_cons
- > (?:l1@〈grid,false〉::l2@〈c1,false〉::l3@〈comma,false〉::rs32@〈bar,false〉::〈d2,true〉::rs3'@〈grid,false〉::rs
- = (l1@〈grid,false〉::l2@〈c1,false〉::l3@〈comma,false〉::rs32@[〈bar,false〉])@〈d2,true〉::rs3'@〈grid,false〉::rs)
- [|>associative_append normalize
- >associative_append normalize
- >associative_append normalize
- >associative_append normalize
- % ]
- >reverse_append >reverse_append >reverse_cons
- >reverse_reverse >reverse_cons >reverse_reverse
- >reverse_append >reverse_append >reverse_cons
- >reverse_reverse >reverse_reverse >reverse_reverse
- >(?:(la@[〈c',false〉])@((((lb@[〈grid,false〉])@l2)@la)@[〈d',false〉])@rs3
- =((la@〈c',false〉::lb)@([〈grid,false〉]@l2@la@[〈d',false〉]@rs3)))
- [|>associative_append >associative_append
- >associative_append >associative_append >associative_append
- >associative_append % ]
- <H2 normalize in ⊢ (??%?); >Hrs3
- >associative_append >associative_append normalize
- >associative_append >associative_append
- @eq_f @eq_f @eq_f
- >(?:la@(〈d',false〉::lc@〈comma,false〉::rs32)@〈bar,false〉::〈d2,true〉::rs3'@〈grid,false〉::rs0 =
- (la@〈d',false〉::lc)@〈comma,false〉::rs32@〈bar,false〉::〈d2,true〉::rs3'@〈grid,false〉::rs0 )
- [| >associative_append normalize >associative_append % ]
- <H3 %
- ]
- ]
- ]
- |* #Hnobars #Htapee >Htapee -Htapee *
- [whd in ⊢ (%→?); * #tapef * whd in ⊢ (%→?); * #_
- #Htapef >Htapef -Htapef
- whd in ⊢ (%→?); * #Htapeout #_ %2 %
- [% [//] whd #x #Hx @Hnobars @memb_append_l2 @memb_cons //
- | >(Htapeout … (refl …)) % ]
- |whd in ⊢ (%→?); * #tapef * whd in ⊢ (%→?);
- * #Hc0 lapply(Hc0 … (refl … )) normalize in ⊢ (%→?); #Htemp destruct (Htemp)
- ]
- |(* no marks in table *)
- #x #membx @(no_marks_in_table … Htable)
- @memb_append_l2
- cut (∀A,l1,l2.∀a:A. a::l1@l2=(a::l1)@l2) [//] #Hcut >Hcut
- >H3 >associative_append @memb_append_l2 @memb_cons @membx
- |(* no grids in table *)
- #x #membx @(no_grids_in_table … Htable)
- @memb_append_l2
- cut (∀A,l1,l2.∀a:A. a::l1@l2=(a::l1)@l2) [//] #Hcut >Hcut
- >H3 >associative_append @memb_append_l2 @memb_cons @membx
- |whd in ⊢ (??%?); >(bit_or_null_not_grid … Hd') >(bit_or_null_not_bar … Hd') %
- ]
- ]
- |#x #membx @(no_marks_in_table … Htable)
- @memb_append_l2 @memb_cons @memb_append_l1 @membx
- |#x #membx @(no_marks_in_table … Htable)
- @memb_append_l1 @membx
- |%
- ]
- ]
-qed.
-
-(*
- MATCH TUPLE
-
- scrolls through the tuples in the transition table until one matching the
- current configuration is found
-*)
-
-definition match_tuple ≝ whileTM ? match_tuple_step (inr … (inl … (inr … start_nop))).
-
-lemma is_grid_true : ∀c.is_grid c = true → c = grid.
-* normalize [ #b ] #H // destruct (H)
-qed.
-
-(* possible variante ?
-definition weakR_match_tuple ≝ λt1,t2.
- (∀ls,cur,rs,b. t1 = midtape STape ls 〈grid,b〉 rs → t2 = t1) ∧
- (∀c,l1,c1,l2,l3,ls0,rs0,n.
- t1 = midtape STape (〈grid,false〉::ls0) 〈bit c,true〉 rs
- (l1@〈grid,false〉::l2@〈bit c1,true〉::l3@〈grid,false〉::rs0) →
- only_bits_or_nulls l1 → no_marks l1 → S n = |l1| →
- table_TM (S n) (l2@〈c1,false〉::l3) →
- (* facciamo match *)
- (∃l4,newc,mv,l5.
- 〈c1,false〉::l3 = l4@〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l5 ∧
- t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls0) 〈grid,false〉
- (l2@l4@〈c,false〉::l1@〈comma,true〉::newc@〈comma,false〉::mv::l5@
- 〈grid,false〉::rs0))
- ∨
- (* non facciamo match su nessuna tupla;
- non specifichiamo condizioni sul nastro di output, perché
- non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
- (current ? t2 = Some ? 〈grid,true〉 ∧
- ∀l4,newc,mv,l5.
- 〈c1,false〉::l3 ≠ l4@〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l5)).
-*)
-
-definition R_match_tuple0 ≝ λt1,t2.
- ∀ls,cur,rs.
- t1 = midtape STape ls cur rs →
- (is_grid (\fst cur) = true → t2 = t1) ∧
- (∀c,l1,c1,l2,l3,ls0,rs0,n.
- ls = 〈grid,false〉::ls0 →
- cur = 〈c,true〉 →
- rs = l1@〈grid,false〉::l2@〈bar,false〉::〈c1,true〉::l3@〈grid,false〉::rs0 →
- is_bit c = true → is_bit c1 = true →
- only_bits_or_nulls l1 → no_marks l1 → S n = |l1| →
- table_TM (S n) (l2@〈bar,false〉::〈c1,false〉::l3) →
- (* facciamo match *)
- (∃l4,newc,mv,l5.
- 〈bar,false〉::〈c1,false〉::l3 = l4@〈bar,false〉::〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l5 ∧
- t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls0) 〈grid,false〉
- (l2@l4@〈bar,false〉::〈c,false〉::l1@〈comma,true〉::newc@〈comma,false〉::mv::l5@
- 〈grid,false〉::rs0))
- ∨
- (* non facciamo match su nessuna tupla;
- non specifichiamo condizioni sul nastro di output, perché
- non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
- (current ? t2 = Some ? 〈grid,true〉 ∧
- ∀l4,newc,mv,l5.
- 〈bar,false〉::〈c1,false〉::l3 ≠ l4@〈bar,false〉::〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l5)).
-
-axiom table_bit_after_bar :
- ∀n,l1,c,l2.table_TM n (l1@〈bar,false〉::〈c,false〉::l2) → is_bit c = true.
-
-lemma wsem_match_tuple : WRealize ? match_tuple R_match_tuple0.
-#intape #k #outc #Hloop
-lapply (sem_while … sem_match_tuple_step intape k outc Hloop) [%] -Hloop
-* #ta * #Hstar @(star_ind_l ??????? Hstar)
-[ whd in ⊢ (%→?); #Hleft
- #ls #cur #rs #Htb cases (Hleft … Htb) #Hgrid #Houtc %
- [ #_ @Houtc
- | #c #l1 #c1 #l2 #l3 #ls0 #rs0 #n #Hls #Hcur #Hrs
- >Hcur in Hgrid; #Hgrid >(is_grid_true … Hgrid) normalize in ⊢ (%→?);
- #Hc destruct (Hc)
- ]
-| (* in the interesting case, we execute a true iteration, then we restart the
- while cycle, finally we end with a false iteration *)
- #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 *)
- #Hcur cases Htc #ls' * #c' * #rs' * * >Htb #Hdes destruct (Hdes)
- #Hfalse @False_ind @(absurd … (is_grid_true … Hcur) Hfalse)
- | (* current and a tuple are marked *)
- #c #l1 #c1 #l2 #l3 #ls0 #rs0 #n #Hls #Hcur #Hrs #Hc #Hc1 #Hl1bitnull #Hl1marks
- #Hl1len #Htable
- cases Htc #ls' * #c' * #rs' * * >Htb #Hdes destruct (Hdes)
- -Htb * #_ #Htc
- (* expose the marked tuple in table *)
- cut (∃la,lb,mv,lc.l3 = la@〈comma,false〉::lb@〈comma,false〉::mv::lc ∧
- S n = |la| ∧ only_bits_or_nulls la)
- [@daemon] * #la * #lb * #mv * #lc * * #Hl3 #Hlalen #Hlabitnull
- >Hl3 in Htable; >append_cons #Htable
- >(?: l2@〈bar,false〉::〈c1,true〉::l3@〈grid,false〉::rs0
- = (l2@[〈bar,false〉])@〈c1,true〉::la@〈comma,false〉::(lb@〈comma,false〉::mv::
- lc)@〈grid,false〉::rs0) in Hrs;
- [| >associative_append normalize >Hl3
- >associative_append normalize % ] #Hrs
- cases (Htc ????????? Hl1bitnull Hl1marks ?? Hlabitnull Hl1len ? Htable Hls Hcur Hrs)
- [5: <Hl1len @Hlalen
- |4: whd in ⊢ (??%?); >Hc1 %
- |3: whd in ⊢ (??%?); >Hc %
- |-Htc *
- [ (* case 1: match successful *)
- * #Heq #Htc % %{[]} %{lb} %{mv} %{lc} destruct (Heq) %
- [%
- | cases (IH … Htc) -IH #Houtc #_ >(Houtc (refl ??))
- >Htc @eq_f normalize >associative_append normalize
- >associative_append normalize %
- ]
- | (* case 2: tuples don't match, we still have other tuples to try *)
- * #Hdiff * #c2 * #l5 * #l6 * #Heqlblc #Htc
- cases (IH ??? … Htc) -IH #_ #IH
- (* by induction hypothesis *)
- lapply (IH ? l1 c2 (l2@〈bar,false〉::〈c1,false〉::la@〈comma,false〉::l5) l6 ? rs0 n (refl ??) (refl ??) ???????)
- [ generalize in match Htable;
- >associative_append normalize
- >associative_append normalize >Heqlblc
- >associative_append normalize //
- | @Hl1len
- | @Hl1marks
- | @Hl1bitnull
- | (*???*) @daemon
- | @Hc
- | >associative_append normalize
- >associative_append normalize
- >associative_append %
- |-IH *
- [ (* the while finally matches a tuple *)
- * #l7 * #newc * #mv0 * #l8 * #Hl7l8 #Houtc %
- >Heqlblc @(ex_intro ?? (〈bar,false〉::〈c1,false〉::la@〈comma,false〉::l5@l7))
- %{newc} %{mv0} %{l8} %
- [ normalize >Hl7l8 >associative_append normalize
- >associative_append %
- | >Houtc @eq_f >associative_append normalize
- >associative_append normalize >associative_append
- normalize >associative_append %
- ]
- | (* the while fails finding a tuple: there are no matches in the whole table *)
- * #Houtc #Hdiff1 %2 %
- [ @Houtc
- | #l50 #newc #mv0 #l51 >Heqlblc
- @daemon
- ]
- ]
- ]
- ]
- | (* 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 ⊢ (???% → %);
- [ #_ @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.
-
-(* termination *)
-lemma WF_mts_niltape:
- WF ? (inv ? R_match_tuple_step_true) (niltape (FinProd FSUnialpha FinBool)).
-@wf #t1 whd in ⊢ (%→?); * #ls * #c * #rs * * #H destruct
-qed.
-
-lemma WF_mts_rightof:
- ∀a,ls. WF ? (inv ? R_match_tuple_step_true) (rightof (FinProd FSUnialpha FinBool) a ls).
-#a #ls @wf #t1 whd in ⊢ (%→?); * #ls * #c * #rs * * #H destruct
-qed.
-
-lemma WF_mts_leftof:
- ∀a,ls. WF ? (inv ? R_match_tuple_step_true) (leftof (FinProd FSUnialpha FinBool) a ls).
-#a #ls @wf #t1 whd in ⊢ (%→?); * #ls * #c * #rs * * #H destruct
-qed.
-
-lemma WF_cst_midtape_grid:
- ∀ls,b,rs. WF ? (inv ? R_match_tuple_step_true)
- (midtape (FinProd … FSUnialpha FinBool) ls 〈grid,b〉 rs).
-#ls #b #rs @wf #t1 whd in ⊢ (%→?); * #ls' * #c' * #rs' * * #H destruct
-* #Hfalse @False_ind @Hfalse %
-qed.
-
-definition Pre_match_tuple ≝ λt.
- ∃ls,cur,rs. t = midtape STape ls cur rs ∧
- (is_grid (\fst cur) = true ∨
- (∃ls0,c,l1,l2,c1,l3,l4,rs0,n.
- only_bits_or_nulls l1 ∧ no_marks l1 ∧
- bit_or_null c = true ∧ bit_or_null c1 = true ∧
- only_bits_or_nulls l3 ∧ S n = |l1| ∧|l1| = |l3| ∧
- table_TM (S n) (l2@〈c1,false〉::l3@〈comma,false〉::l4) ∧
- ls = 〈grid,false〉::ls0 ∧ cur = 〈c,true〉 ∧
- rs = l1@〈grid,false〉::l2@〈c1,true〉::l3@〈comma,false〉::l4@〈grid,false〉::rs0)).
-
-lemma acc_Realize_to_acc_GRealize: ∀sig,M.∀q:states sig M.∀P,R1,R2.
- M ⊨ [q:R1,R2] → accGRealize sig M q P R1 R2.
-#alpha #M #q #Pre #R1 #R2 #HR #t #HPre
-cases (HR t) -HR #k * #outc * * #Hloop #HRtrue #HRfalse
-@(ex_intro ?? k) @(ex_intro ?? outc) %
- [ % [@Hloop] @HRtrue | @HRfalse]
-qed.
-
-
-lemma terminate_match_tuple:
- ∀t. Pre_match_tuple t → Terminate ? match_tuple t.
-#t #HPre
-@(terminate_while_guarded ???
- Pre_match_tuple …
- (acc_Realize_to_acc_GRealize ??? Pre_match_tuple … sem_match_tuple_step)
- … HPre) [%]
- [-HPre -t #t1 #t2 #HPre cases HPre #ls * * #curl #curr * #rs * #Ht1 *
- [(* absurd case *)
- #Hgrid * #ls1 * #cur1 * #rs1 * * >Ht1 #Hdes destruct (Hdes)
- #Habs @False_ind @(absurd ?? Habs) @(is_grid_true … Hgrid)
- |* #ls0 * #c * #l1 * #l2 * #c1 * #l3 * #l4 * #rs0 * #n
- * * * * * * * * * *
- #Hl1 #Hmarksl1 #Hc #Hc1 #Hl3 #lenl1 #eqlen #Htable #Hls #Hcur #Hrs
- * #ls1 * #cur1 * #rs1 * * >Ht1 #Hdes destruct (Hdes) #Hdes #H
- lapply (H … Hl1 Hmarksl1 Hc Hc1 Hl3 lenl1 eqlen Htable Hls Hcur Hrs)
- -H *
- [* [ * #Hdes #Ht2 >Ht2
- @ex_intro [2:@ex_intro [2: @ex_intro [2: % [%]|]|]|]
- %1 %
- |* #test * #c2 * #l5 * #l6 * #Hl4 #Ht2
- cut (∃l7,l8. l6 = l7@〈comma,false 〉::l8 ∧ |l7| = |l1|) [@daemon]
- * #l7 * #l8 * #Hl6 #eqlen1
- @ex_intro [2:@ex_intro [2: @ex_intro [2: % [@Ht2]|]|]|] %2
- @(ex_intro … ls0) @(ex_intro … c) @(ex_intro … l1)
- @(ex_intro … (l2@〈c1,false〉::l3@〈comma,false〉::l5@[〈bar,false〉]))
- @(ex_intro … c2) @(ex_intro … l7) @(ex_intro … l8)
- @(ex_intro … rs0) @(ex_intro … n)
- % [2: >Hl6 >associative_append >associative_append @eq_f @eq_f @eq_f
- @eq_f >associative_append @eq_f @eq_f >associative_append % ]
- % [2: %] % [2: %] % [2:@daemon] % [2: @sym_eq @eqlen1]
- % [2: @lenl1] % [2: #x #memx @daemon]
- % [2: @daemon] % [2: @Hc] % [2: @Hmarksl1] @Hl1
- ]
- |* * #_ #_ #H cases (current_to_midtape … H) #ls * #rs #Ht1
- >Ht1 @ex_intro [2:@ex_intro [2: @ex_intro [2: % [%]|]|]|] %1 %
- ]
- ]
- |cases HPre -HPre #ls * * #curl #curr * #rs * #Ht *
- [#Hgrid >Ht >(is_grid_true … Hgrid) @WF_cst_midtape_grid
- |* #ls0 * #c * #l1 * #l2 * #c1 * #l3 * #l4
- cut (∃len. |l4| = len) [/2/] * #lenl4
- lapply l4 lapply l3 lapply c1 lapply l2 lapply l1 lapply c lapply ls0 lapply Ht
- lapply curr lapply curl lapply ls lapply rs lapply t -l4 -l3 -l2 -l1 -c1 -curr -curl -ls -t
- -c -ls0 -rs
- (* by induction on the length of l4 *)
- @(nat_elim1 lenl4)
- #len #Hind #t #rs #ls #cl #cr #Ht #ls0 #c #l1 #l2 #c1 #l3 #l4 #Hlen
- * #rs0 * #n * * * * * * * * * *
- #Hl1 #Hmarksl1 #Hc #Hc1 #Hl3 #lenl1 #eqlen #Htable #Hls #Hcur #Hrs
- % #t1 >Ht whd in ⊢ (%→?); * #ls1 * #cur * #rs1 * * #Hdes destruct (Hdes)
- #Hgrid #H lapply (H … Hl1 Hmarksl1 Hc Hc1 Hl3 lenl1 eqlen Htable Hls Hcur Hrs)
- -H *
- [* [ * #Hdes destruct (Hdes) #Ht1 >Ht1 @WF_cst_midtape_grid
- | * #_ * #c2 * #l5 * #l6 * #Hl4 #Ht1
- cut (∃l7,l8. l6 = l7@〈comma,false 〉::l8 ∧ |l7| = |l1|) [@daemon]
- * #l7 * #l8 * #Hl6 #eqlen1
- @(Hind … Ht1 ls0 c l1 (l2@〈c1,false〉::l3@〈comma,false〉::l5@[〈bar,false〉]) c2 l7 l8 … (refl …))
- [<Hlen >Hl4 >Hl6 >length_append normalize in match (length … (cons …));
- >length_append normalize in match (length … (cons …)); <plus_n_Sm
- @le_S_S @daemon
- |@(ex_intro … rs0) @(ex_intro … n) %
- [2: >Hl6 >associative_append >associative_append @eq_f @eq_f @eq_f
- @eq_f >associative_append @eq_f @eq_f >associative_append % ]
- % [2: %] % [2: %] % [2:@daemon] % [2: @sym_eq @eqlen1]
- % [2: @lenl1] % [2: #x #memx @daemon]
- % [2: @daemon] % [2: @Hc] % [2: @Hmarksl1] @Hl1
- ]
- ]
- |* * #_ #_ #H cases (current_to_midtape … H) #ls * #rs #Ht1
- >Ht1 //
- ]
- ]
-qed.
-
-definition R_match_tuple ≝ λt1,t2.
- ∀ls,c,l1,c1,l2,rs,n.
- is_bit c = true → is_bit c1 = true →
- only_bits_or_nulls l1 → no_marks l1 → S n = |l1| →
- table_TM (S n) (〈bar,false〉::〈c1,false〉::l2) →
- t1 = midtape STape (〈grid,false〉::ls) 〈c,true〉
- (l1@〈grid,false〉::〈bar,false〉::〈c1,true〉::l2@〈grid,false〉::rs) →
- (* facciamo match *)
- (∃l3,newc,mv,l4.
- 〈bar,false〉::〈c1,false〉::l2 = l3@〈bar,false〉::〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l4 ∧
- t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls) 〈grid,false〉
- (l3@〈bar,false〉::〈c,false〉::l1@〈comma,true〉::newc@〈comma,false〉::mv::l4@〈grid,false〉::rs))
- ∨
- (* non facciamo match su nessuna tupla;
- non specifichiamo condizioni sul nastro di output, perché
- non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
- (current ? t2 = Some ? 〈grid,true〉 ∧
- ∀l3,newc,mv,l4.
- 〈bar,false〉::〈c1,false〉::l2 ≠ l3@〈bar,false〉::〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l4).
-
-lemma sem_match_tuple0 : GRealize ? match_tuple Pre_match_tuple R_match_tuple0.
-@WRealize_to_GRealize [@terminate_match_tuple | @wsem_match_tuple]
-qed.
-
-lemma sem_match_tuple : GRealize ? match_tuple Pre_match_tuple R_match_tuple.
-generalize in match sem_match_tuple0; @GRealize_to_GRealize
-#t1 #t2 #HR #ls #c #l1 #c1 #l2 #rs #n #Hc #Hc1 #Hl1bitsnulls #Hl1marks #Hl1len #Htable #Ht1
-cases (HR … Ht1) -HR #_ #HR
-@(HR ??? [] … (refl ??) (refl ??) (refl ??) Hc Hc1 Hl1bitsnulls Hl1marks
- Hl1len Htable)
-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/move_char.ma".
-include "turing/universal/marks.ma".
-include "turing/universal/tuples.ma".
-
-definition init_cell_states ≝ initN 2.
-
-definition ics0 : init_cell_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)).
-definition ics1 : init_cell_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)).
-
-definition init_cell ≝
- mk_TM STape init_cell_states
- (λp.let 〈q,a〉 ≝ p in
- match pi1 … q with
- [ O ⇒ match a with
- [ None ⇒ 〈ics1, Some ? 〈〈null,false〉,N〉〉
- | Some _ ⇒ 〈ics1, None ?〉 ]
- | S _ ⇒ 〈ics1,None ?〉 ])
- ics0 (λq.q == ics1).
-
-definition R_init_cell ≝ λt1,t2.
- (∃c.current STape t1 = Some ? c ∧ t2 = t1) ∨
- (current STape t1 = None ? ∧ t2 = midtape ? (left ? t1) 〈null,false〉 (right ? t1)).
-
-axiom sem_init_cell : Realize ? init_cell R_init_cell.
-
-(*
-MOVE TAPE RIGHT:
-
- ls # current c # table # d? rs
- ^
- ls # current c # table # d? rs init
- ^
- ls # current c # table # d? rs
- ^
- ls # current c # table # d rs ----------------------
- ^ move_l
- ls # current c # table # d rs
- ^ swap
- ls # current c # table d # rs --------------------
- ^
- ls # current c # table d # rs
- ^
- ls # current c # d table # rs sub1
- ^
- ls # current c # d table # rs
- ^
- ls # current c d # table # rs -------------------
- ^ move_l
- ls # current c d # table # rs -------------------
- ^
- ls # current c d # table # rs
- ^
- ls # c current d # table # rs sub1
- ^
- ls # c current d # table # rs
- ^
- ls c # current d # table # rs ------------------
- ^
-
-(move_to_grid_r;)
-move_r;
-init_cell;
-move_l;
-swap;
-
-move_l;
-move_char_l;
----------move_l;
-swap;
-
-move_l;
-
-move_l;
-move_char_l;
----------move_l;
-swap
-*)
-
-(* l1 # l2 r ---> l1 r # l2
- ^ ^
- *)
-definition move_after_left_bar ≝
- move_l … · move_char_l STape 〈grid,false〉 · swap_r STape 〈grid,false〉.
-
-definition R_move_after_left_bar ≝ λt1,t2.
- ∀l1,l2,l3,r. t1 = midtape STape (l2@〈grid,false〉::l1) r l3 → no_grids l2 →
- t2 = midtape STape l1 r (〈grid,false〉::reverse ? l2@l3).
-
-lemma sem_move_after_left_bar : Realize ? move_after_left_bar R_move_after_left_bar.
-#intape
-cases (sem_seq … (sem_move_l …) (sem_seq … (ssem_move_char_l STape 〈grid,false〉)
- (sem_swap_r STape 〈grid,false〉)) intape)
-#k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#l1 #l2 #l3 #r #Hintape #Hl2
-cases HR -HR #ta * whd in ⊢ (%→?); #Hta lapply (proj2 ?? Hta … Hintape) -Hta #Hta
-* #tb * whd in ⊢(%→?); generalize in match Hta; -Hta cases l2 in Hl2;
-[ #_ #Hta #Htb lapply (Htb … Hta) -Htb * #Htb lapply (Htb (refl ??)) -Htb #Htb #_
- whd in ⊢(%→?); >Htb #Houtc lapply (proj2 ?? Houtc … Hta) -Houtc #Houtc @Houtc
-| #c0 #l0 #Hnogrids #Hta #Htb lapply (Htb … Hta) -Htb * #_ #Htb
- lapply (Htb … (refl ??) ??)
- [ cases (true_or_false (memb STape 〈grid,false〉 l0)) #Hmemb
- [ @False_ind lapply (Hnogrids 〈grid,false〉 ?)
- [ @memb_cons // | normalize #Hfalse destruct (Hfalse) ]
- | @Hmemb ]
- | % #Hc0 lapply (Hnogrids c0 ?)
- [ @memb_hd | >Hc0 normalize #Hfalse destruct (Hfalse) ]
- | #Htb whd in ⊢(%→?); >Htb #Houtc lapply (proj2 ?? Houtc … (refl ??)) -Houtc #Houtc
- >reverse_cons >associative_append @Houtc
-]]
-qed.
-
-definition move_tape_r ≝
- move_r … · init_cell · move_l … · swap_r STape 〈grid,false〉 ·
- move_after_left_bar · move_l … · move_after_left_bar · move_r ….
-
-definition R_move_tape_r ≝ λt1,t2.
- ∀rs,n,table,c0,bc0,curconfig,ls0.
- bit_or_null c0 = true → only_bits_or_nulls curconfig → table_TM n (reverse ? table) →
- t1 = midtape STape (table@〈grid,false〉::〈c0,bc0〉::curconfig@〈grid,false〉::ls0)
- 〈grid,false〉 rs →
- (rs = [] ∧
- t2 = midtape STape (〈c0,bc0〉::ls0) 〈grid,false〉 (reverse STape curconfig@〈null,false〉::
- 〈grid,false〉::reverse STape table@[〈grid,false〉])) ∨
- (∃r0,rs0.rs = r0::rs0 ∧
- t2 = midtape STape (〈c0,bc0〉::ls0) 〈grid,false〉 (reverse STape curconfig@r0::
- 〈grid,false〉::reverse STape table@〈grid,false〉::rs0)).
-
-lemma sem_move_tape_r : Realize ? move_tape_r R_move_tape_r.
-#tapein
-cases (sem_seq …(sem_move_r …) (sem_seq … sem_init_cell (sem_seq … (sem_move_l …)
- (sem_seq … (sem_swap_r STape 〈grid,false〉) (sem_seq … sem_move_after_left_bar
- (sem_seq … (sem_move_l …) (sem_seq … sem_move_after_left_bar (sem_move_r …))))))) tapein)
-#k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#rs #n #table #c0 #bc0 #curconfig #ls0 #Hbitnullc0 #Hbitnullcc #Htable #Htapein
-cases HR -HR #ta * whd in ⊢ (%→?); #Hta lapply (proj2 ?? Hta … Htapein) -Hta #Hta
-* #tb * whd in ⊢ (%→?); *
-[ * #r0 *
- generalize in match Hta; generalize in match Htapein; -Htapein -Hta cases rs
- [ #_ #Hta >Hta normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
- #r1 #rs1 #Htapein #Hta change with (midtape ?? r1 rs1) in Hta:(???%); >Hta
- #Hr0 whd in Hr0:(??%?); #Htb * #tc * whd in ⊢ (%→?); #Htc lapply (proj2 ?? Htc … Htb) -Htc #Htc
- * #td * whd in ⊢ (%→?); #Htd lapply (proj2 ?? Htd … Htc) -Htd #Htd
- * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd ?) [ #x #memx @(no_grids_in_table … Htable) @memb_reverse @memx]
- -Hte #Hte
- * #tf * whd in ⊢ (%→?); #Htf lapply (proj2 ?? Htf … Hte) -Htf #Htf
- * #tg * whd in ⊢ (%→?); #Htg lapply (Htg … Htf ?) [ #x #Hx @bit_or_null_not_grid @Hbitnullcc // ] -Htg #Htg
- whd in ⊢ (%→?); #Houtc lapply (proj2 ?? Houtc … Htg) -Houtc #Houtc
- %2 @(ex_intro ?? r1) @(ex_intro ?? rs1) % [%] @Houtc
-| * generalize in match Hta; generalize in match Htapein; -Htapein -Hta cases rs
- [|#r1 #rs1 #_ #Hta >Hta normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
- #Htapein #Hta change with (rightof ???) in Hta:(???%); >Hta
- #Hr0 whd in Hr0:(??%?); #Htb * #tc * whd in ⊢ (%→?); #Htc lapply (proj2 ?? Htc … Htb) -Htc #Htc
- * #td * whd in ⊢ (%→?); #Htd lapply (proj2 ?? Htd … Htc) -Htd #Htd
- * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd ?) [ #x #memx @(no_grids_in_table … Htable) @memb_reverse @memx] -Hte #Hte
- * #tf * whd in ⊢ (%→?); #Htf lapply (proj2 ?? Htf … Hte) -Htf #Htf
- * #tg * whd in ⊢ (%→?); #Htg lapply (Htg … Htf ?) [ #x #Hx @bit_or_null_not_grid @Hbitnullcc // ] -Htg #Htg
- whd in ⊢ (%→?); #Houtc lapply (proj2 ?? Houtc … Htg) -Houtc #Houtc
- % % [% | @Houtc ]
-qed.
-
-(*
-MOVE TAPE LEFT:
-
- ls d? # current c # table # rs
- ^ move_l; adv_to_mark_l
- ls d? # current c # table # rs
- ^ move_l; adv_to_mark_l
- ls d? # current c # table # rs
- ^ move_l
- ls d? # current c # table # rs
- ^ init_cell
- ls d # current c # table # rs
- ^ mtl_aux
- ls # current c d # table # rs
- ^ swap_r
- ls # current d c # table # rs
- ^ mtl_aux
- ls # current d # table c # rs
- ^ swap
- ls # current d # table # c rs
- ^ move_l; adv_to_mark_l
- ls # current d # table # c rs
- ^ move_l; adv_to_mark_l
- ls # current d # table # c rs
- ^
-*)
-definition mtl_aux ≝
- swap_r STape 〈grid,false〉 · move_r … · move_r … ·
- move_char_r STape 〈grid,false〉 · move_l ….
-definition R_mtl_aux ≝ λt1,t2.
- ∀l1,l2,l3,r. t1 = midtape STape l1 r (〈grid,false〉::l2@〈grid,false〉::l3) → no_grids l2 →
- t2 = midtape STape (reverse ? l2@〈grid,false〉::l1) r (〈grid,false〉::l3).
-
-lemma sem_mtl_aux : Realize ? mtl_aux R_mtl_aux.
-#intape
-cases (sem_seq … (sem_swap_r STape 〈grid,false〉) (sem_seq … (sem_move_r …)
- (sem_seq … (sem_move_r …) (sem_seq … (ssem_move_char_r STape 〈grid,false〉)
- (sem_move_l …)))) intape)
-#k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#l1 #l2 #l3 #r #Hintape #Hl2
-cases HR -HR #ta * whd in ⊢ (%→?); #Hta lapply (proj2 ?? Hta … Hintape) -Hta #Hta
-* #tb * whd in ⊢(%→?); #Htb lapply (proj2 ?? Htb … Hta) -Htb -Hta whd in ⊢ (???%→?); #Htb
-* #tc * whd in ⊢(%→?); #Htc lapply (proj2 ?? Htc … Htb) -Htc -Htb cases l2 in Hl2;
-[ #_ #Htc * #td * whd in ⊢(%→?); #Htd lapply (Htd … Htc) -Htd >Htc -Htc * #Htd #_
- whd in ⊢ (%→?); #Houtc lapply (Htd (refl ??)) -Htd @(proj2 ?? Houtc)
-| #c0 #l0 #Hnogrids #Htc *
- #td * whd in ⊢(%→?); #Htd lapply (Htd … Htc) -Htd -Htc * #_ #Htd
- lapply (Htd … (refl ??) ??)
- [ cases (true_or_false (memb STape 〈grid,false〉 l0)) #Hmemb
- [ @False_ind lapply (Hnogrids 〈grid,false〉 ?)
- [ @memb_cons // | normalize #Hfalse destruct (Hfalse) ]
- | @Hmemb ]
- | % #Hc0 lapply (Hnogrids c0 ?)
- [ @memb_hd | >Hc0 normalize #Hfalse destruct (Hfalse) ]
- | #Htd whd in ⊢(%→?); >Htd #Houtc lapply (proj2 ?? Houtc … (refl ??)) -Houtc #Houtc
- >reverse_cons >associative_append @Houtc
-]]
-qed.
-
-definition R_ml_atml ≝ λt1,t2.
- ∀ls1,ls2,rs.no_grids ls1 →
- t1 = midtape STape (ls1@〈grid,false〉::ls2) 〈grid,false〉 rs →
- t2 = midtape STape ls2 〈grid,false〉 (reverse ? ls1@〈grid,false〉::rs).
-
-lemma sem_ml_atml :
- Realize ? ((move_l …) · (adv_to_mark_l … (λc:STape.is_grid (\fst c)))) R_ml_atml.
-#intape
-cases (sem_seq … (sem_move_l …) (sem_adv_to_mark_l … (λc:STape.is_grid (\fst c))) intape)
-#k * #outc * #Hloop #HR %{k} %{outc} % [@Hloop] -Hloop
-#ls1 #ls2 #rs #Hnogrids #Hintape cases HR -HR
-#ta * whd in ⊢ (%→?); #Hta lapply (proj2 ?? Hta … Hintape) -Hta
-cases ls1 in Hnogrids;
-[ #_ #Hta whd in ⊢ (%→?); #Houtc cases (proj2 ?? Houtc … Hta) -Houtc
- #Houtc #_ >Houtc [@Hta | %]
-| #c0 #l0 #Hnogrids #Hta whd in ⊢ (%→?); #Houtc cases (proj2 ?? Houtc … Hta) -Houtc
- #_ #Houtc cases (Houtc ?)
- [#Houtc #_ >(Houtc … (refl …)(refl …))
- [>reverse_cons >associative_append %
- | #x #Hx @Hnogrids @memb_cons //
- ]
- |@Hnogrids @memb_hd
- ]
-]
-qed.
-
-definition move_tape_l : TM STape ≝
- (move_l … · adv_to_mark_l … (λc:STape.is_grid (\fst c))) ·
- (move_l … · adv_to_mark_l … (λc:STape.is_grid (\fst c))) ·
- move_l … · init_cell · mtl_aux · swap_l STape 〈grid,false〉 ·
- mtl_aux ·swap_r STape 〈grid,false〉 ·
- (move_l … · adv_to_mark_l … (λc:STape.is_grid (\fst c))) ·
- (move_l … · adv_to_mark_l … (λc:STape.is_grid (\fst c))).
-
-(* seq ? (move_r …) (seq ? init_cell (seq ? (move_l …)
- (seq ? (swap STape 〈grid,false〉)
- (seq ? move_after_left_bar (seq ? (move_l …) move_after_left_bar))))). *)
-
-axiom daemon: ∀P:Prop.P.
-
-definition R_move_tape_l ≝ λt1,t2.
- ∀rs,n,table,c0,bc0,curconfig,ls0.
- bit_or_null c0 = true → only_bits_or_nulls curconfig →
- table_TM n (reverse ? table) → only_bits ls0 →
- t1 = midtape STape (table@〈grid,false〉::〈c0,bc0〉::curconfig@〈grid,false〉::ls0)
- 〈grid,false〉 rs →
- (ls0 = [] ∧
- t2 = midtape STape [] 〈grid,false〉
- (reverse ? curconfig@〈null,false〉::〈grid,false〉::reverse ? table@〈grid,false〉::〈c0,bc0〉::rs)) ∨
- (∃l1,ls1. ls0 = l1::ls1 ∧
- t2 = midtape STape ls1 〈grid,false〉
- (reverse ? curconfig@l1::〈grid,false〉::reverse ? table@〈grid,false〉::〈c0,bc0〉::rs)).
-
-lemma sem_move_tape_l : Realize ? move_tape_l R_move_tape_l.
-#tapein
-cases (sem_seq … sem_ml_atml
- (sem_seq … sem_ml_atml
- (sem_seq … (sem_move_l …)
- (sem_seq … sem_init_cell
- (sem_seq … sem_mtl_aux
- (sem_seq … (sem_swap_l STape 〈grid,false〉)
- (sem_seq … sem_mtl_aux
- (sem_seq … (sem_swap_r STape 〈grid,false〉)
- (sem_seq … sem_ml_atml sem_ml_atml)))))))) tapein)
-#k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#rs #n #table #c0 #bc0 #curconfig #ls0 #Hbitnullc0 #Hbitnullcc #Htable #Hls0 #Htapein
-cases HR -HR #ta * whd in ⊢ (%→?); #Hta lapply (Hta … Htapein)
-[#x #memx @(no_grids_in_table … Htable) @memb_reverse @memx]
--Hta #Hta * #tb * whd in ⊢ (%→?); #Htb lapply (Htb (〈c0,bc0〉::curconfig) … Hta)
-[ @daemon ] -Hta -Htb #Htb
-* #tc * whd in ⊢ (%→?); #Htc lapply (proj2 ?? Htc … Htb) -Htb -Htc #Htc
-* #td * whd in ⊢ (%→?); *
-[ * #c1 * generalize in match Htc; generalize in match Htapein; -Htapein -Htc
- cases ls0 in Hls0;
- [ #_ #_ #Htc >Htc normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
- #l1 #ls1 #Hls0 #Htapein #Htc change with (midtape ? ls1 l1 ?) in Htc:(???%); >Htc
- #Hl1 whd in Hl1:(??%?); #Htd
- * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd ?)
- [ (* memb_reverse *) @daemon ] -Hte -Htd >reverse_reverse #Hte
- * #tf * whd in ⊢ (%→?); #Htf lapply (proj2 ?? Htf … Hte) -Htf -Hte #Htf
- * #tg * whd in ⊢ (%→?); #Htg lapply (Htg … Htf ?)
- [ @(no_grids_in_table … Htable) ] -Htg -Htf >reverse_reverse #Htg
- * #th * whd in ⊢ (%→?); #Hth lapply (proj2 ?? Hth … Htg) -Hth -Htg #Hth
- * #ti * whd in ⊢ (%→?); #Hti lapply (Hti … Hth)
- [ (* memb_reverse *) @daemon ] -Hti -Hth #Hti
- whd in ⊢ (%→?); #Houtc lapply (Houtc (l1::curconfig) … Hti)
- [ #x #Hx cases (orb_true_l … Hx) -Hx #Hx
- [ >(\P Hx) lapply (Hls0 l1 (memb_hd …)) @bit_not_grid
- | lapply (Hbitnullcc ? Hx) @bit_or_null_not_grid ] ]
- -Houtc >reverse_cons >associative_append #Houtc %2 %{l1} %{ls1} % [%] @Houtc
-| * generalize in match Htc; generalize in match Htapein; -Htapein -Htc
- cases ls0
- [| #l1 #ls1 #_ #Htc >Htc normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
- #Htapein #Htc change with (leftof ???) in Htc:(???%); >Htc #_ #Htd
- * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd ?)
- [ (*memb_reverse*) @daemon ] -Hte -Htd >reverse_reverse #Hte
- * #tf * whd in ⊢ (%→?); #Htf lapply (proj2 ?? Htf … Hte) -Htf -Hte #Htf
- * #tg * whd in ⊢ (%→?); #Htg lapply (Htg … Htf ?)
- [ @(no_grids_in_table … Htable) ] -Htg -Htf >reverse_reverse #Htg
- * #th * whd in ⊢ (%→?); #Hth lapply (proj2 ?? Hth … Htg) -Hth -Htg #Hth
- * #ti * whd in ⊢ (%→?); #Hti lapply (Hti … Hth)
- [ (*memb_reverse*) @daemon ] -Hti -Hth #Hti
- whd in ⊢ (%→?); #Houtc lapply (Houtc (〈null,false〉::curconfig) … Hti)
- [ #x #Hx cases (orb_true_l … Hx) -Hx #Hx
- [ >(\P Hx) %
- | lapply (Hbitnullcc ? Hx) @bit_or_null_not_grid ] ]
- -Houtc >reverse_cons >associative_append
- >reverse_cons >associative_append #Houtc % % [%] @Houtc
-]
-qed.
-
-(*definition mtl_aux ≝
- seq ? (move_r …) (seq ? (move_char_r STape 〈grid,false〉) (move_l …)).
-definition R_mtl_aux ≝ λt1,t2.
- ∀l1,l2,l3,r. t1 = midtape STape l1 r (l2@〈grid,false〉::l3) → no_grids l2 →
- t2 = midtape STape (reverse ? l2@l1) r (〈grid,false〉::l3).
-
-lemma sem_mtl_aux : Realize ? mtl_aux R_mtl_aux.
-#intape
-cases (sem_seq … (sem_move_r …) (sem_seq … (ssem_move_char_r STape 〈grid,false〉) (sem_move_l …)) intape)
-#k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#l1 #l2 #l3 #r #Hintape #Hl2
-cases HR -HR #ta * whd in ⊢ (%→?); #Hta lapply (Hta … Hintape) -Hta #Hta
-* #tb * whd in ⊢(%→?); generalize in match Hta; -Hta cases l2 in Hl2;
-[ #_ #Hta #Htb lapply (Htb … Hta) -Htb * #Htb #_ whd in ⊢ (%→?); #Houtc
- lapply (Htb (refl ??)) -Htb >Hta @Houtc
-| #c0 #l0 #Hnogrids #Hta #Htb lapply (Htb … Hta) -Htb * #_ #Htb
- lapply (Htb … (refl ??) ??)
- [ cases (true_or_false (memb STape 〈grid,false〉 l0)) #Hmemb
- [ @False_ind lapply (Hnogrids 〈grid,false〉 ?)
- [ @memb_cons // | normalize #Hfalse destruct (Hfalse) ]
- | @Hmemb ]
- | % #Hc0 lapply (Hnogrids c0 ?)
- [ @memb_hd | >Hc0 normalize #Hfalse destruct (Hfalse) ]
- | #Htb whd in ⊢(%→?); >Htb #Houtc lapply (Houtc … (refl ??)) -Houtc #Houtc
- >reverse_cons >associative_append @Houtc
-]]
-qed.
-
-check swap*)
-
-
-(*
- by cases on current:
- case bit false: move_tape_l
- case bit true: move_tape_r
- case null: adv_to_grid_l; move_l; adv_to_grid_l;
-*)
-
-definition lift_tape ≝ λls,c,rs.
- let 〈c0,b〉 ≝ c in
- let c' ≝ match c0 with
- [ null ⇒ None ?
- | _ ⇒ Some ? c ]
- in
- mk_tape STape ls c' rs.
-
-definition sim_current_of_tape ≝ λt.
- match current STape t with
- [ None ⇒ 〈null,false〉
- | Some c0 ⇒ c0 ].
-
-
-definition move_of_unialpha ≝
- λc.match c with
- [ bit x ⇒ match x with [ true ⇒ R | false ⇒ L ]
- | _ ⇒ N ].
-
-definition no_nulls ≝
- λl:list STape.∀x.memb ? x l = true → is_null (\fst x) = false.
-
-definition current_of_alpha ≝ λc:STape.
- match \fst c with [ null ⇒ None ? | _ ⇒ Some ? c ].
-
-(*
- no_marks (c::ls@rs)
- only_bits (ls@rs)
- bit_or_null c
-
-*)
-definition legal_tape ≝ λls,c,rs.
- no_marks (c::ls@rs) ∧ only_bits (ls@rs) ∧ bit_or_null (\fst c) = true ∧
- (\fst c ≠ null ∨ ls = [] ∨ rs = []).
-
-lemma legal_tape_left :
- ∀ls,c,rs.legal_tape ls c rs →
- left ? (mk_tape STape ls (current_of_alpha c) rs) = ls.
-#ls * #c #bc #rs * * * #_ #_ #_ *
-[ *
- [ cases c
- [ #c' #_ %
- | * #Hfalse @False_ind /2/
- |*: #_ % ]
- | #Hls >Hls cases c // cases rs //
- ]
-| #Hrs >Hrs cases c // cases ls //
-]
-qed.
-
-axiom legal_tape_current :
- ∀ls,c,rs.legal_tape ls c rs →
- current ? (mk_tape STape ls (current_of_alpha c) rs) = current_of_alpha c.
-
-axiom legal_tape_right :
- ∀ls,c,rs.legal_tape ls c rs →
- right ? (mk_tape STape ls (current_of_alpha c) rs) = rs.
-
-(*
-lemma legal_tape_cases :
- ∀ls,c,rs.legal_tape ls c rs →
- \fst c ≠ null ∨ (\fst c = null ∧ (ls = [] ∨ rs = [])).
-#ls #c #rs cases c #c0 #bc0 cases c0
-[ #c1 normalize #_ % % #Hfalse destruct (Hfalse)
-| cases ls
- [ #_ %2 % // % %
- | #l0 #ls0 cases rs
- [ #_ %2 % // %2 %
- | #r0 #rs0 normalize * * #_ #Hrs destruct (Hrs) ]
- ]
-|*: #_ % % #Hfalse destruct (Hfalse) ]
-qed.
-
-axiom legal_tape_conditions :
- ∀ls,c,rs.(\fst c ≠ null ∨ ls = [] ∨ rs = []) → legal_tape ls c rs.
-(*#ls #c #rs *
-[ *
- [ >(eq_pair_fst_snd ?? c) cases (\fst c)
- [ #c0 #Hc % % %
- | * #Hfalse @False_ind /2/
- |*: #Hc % % %
- ]
- | cases ls [ * #Hfalse @False_ind /2/ ]
- #l0 #ls0
-
- #Hc
-*)
-*)
-
-definition R_move_tape_r_abstract ≝ λt1,t2.
- ∀rs,n,table,curc,curconfig,ls.
- is_bit curc = true → only_bits_or_nulls curconfig → table_TM n (reverse ? table) →
- t1 = midtape STape (table@〈grid,false〉::〈curc,false〉::curconfig@〈grid,false〉::ls)
- 〈grid,false〉 rs →
- legal_tape ls 〈curc,false〉 rs →
- ∀t1'.t1' = lift_tape ls 〈curc,false〉 rs →
- ∃ls1,rs1,newc.
- (t2 = midtape STape ls1 〈grid,false〉 (reverse ? curconfig@〈newc,false〉::
- 〈grid,false〉::reverse ? table@〈grid,false〉::rs1) ∧
- lift_tape ls1 〈newc,false〉 rs1 =
- tape_move_right STape ls 〈curc,false〉 rs ∧ legal_tape ls1 〈newc,false〉 rs1).
-
-lemma lift_tape_not_null :
- ∀ls,c,rs. is_null (\fst c) = false →
- lift_tape ls c rs = mk_tape STape ls (Some ? c) rs.
-#ls * #c0 #bc0 #rs cases c0
-[|normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
-//
-qed.
-
-axiom bit_not_null : ∀d.is_bit d = true → is_null d = false.
-
-lemma mtr_concrete_to_abstract :
- ∀t1,t2.R_move_tape_r t1 t2 → R_move_tape_r_abstract t1 t2.
-#t1 #t2 whd in ⊢(%→?); #Hconcrete
-#rs #n #table #curc #curconfig #ls #Hbitcurc #Hcurconfig #Htable #Ht1
-* * * #Hnomarks #Hbits #Hcurc #Hlegal #t1' #Ht1'
-cases (Hconcrete … Htable Ht1) //
-[ * #Hrs #Ht2
- @(ex_intro ?? (〈curc,false〉::ls)) @(ex_intro ?? [])
- @(ex_intro ?? null) %
- [ %
- [ >Ht2 %
- | >Hrs % ]
- | % [ % [ %
- [ >append_nil #x #Hx cases (orb_true_l … Hx) #Hx'
- [ >(\P Hx') %
- | @Hnomarks @(memb_append_l1 … Hx') ]
- | >append_nil #x #Hx cases (orb_true_l … Hx) #Hx'
- [ >(\P Hx') //
- | @Hbits @(memb_append_l1 … Hx') ]]
- | % ]
- | %2 % ]
- ]
-| * * #r0 #br0 * #rs0 * #Hrs
- cut (br0 = false)
- [ @(Hnomarks 〈r0,br0〉) @memb_cons @memb_append_l2 >Hrs @memb_hd]
- #Hbr0 >Hbr0 in Hrs; #Hrs #Ht2
- @(ex_intro ?? (〈curc,false〉::ls)) @(ex_intro ?? rs0)
- @(ex_intro ?? r0) %
- [ %
- [ >Ht2 //
- | >Hrs >lift_tape_not_null
- [ %
- | @bit_not_null @(Hbits 〈r0,false〉) >Hrs @memb_append_l2 @memb_hd ] ]
- | % [ % [ %
- [ #x #Hx cases (orb_true_l … Hx) #Hx'
- [ >(\P Hx') %
- | cases (memb_append … Hx') #Hx'' @Hnomarks
- [ @(memb_append_l1 … Hx'')
- | >Hrs @memb_cons @memb_append_l2 @(memb_cons … Hx'') ]
- ]
- | whd in ⊢ (?%); #x #Hx cases (orb_true_l … Hx) #Hx'
- [ >(\P Hx') //
- | cases (memb_append … Hx') #Hx'' @Hbits
- [ @(memb_append_l1 … Hx'') | >Hrs @memb_append_l2 @(memb_cons … Hx'') ]
- ]]
- | whd in ⊢ (??%?); >(Hbits 〈r0,false〉) //
- @memb_append_l2 >Hrs @memb_hd ]
- | % % % #Hr0 lapply (Hbits 〈r0,false〉?)
- [ @memb_append_l2 >Hrs @memb_hd
- | >Hr0 normalize #Hfalse destruct (Hfalse)
- ] ] ] ]
-qed.
-
-definition R_move_tape_l_abstract ≝ λt1,t2.
- ∀rs,n,table,curc,curconfig,ls.
- is_bit curc = true → only_bits_or_nulls curconfig → table_TM n (reverse ? table) →
- t1 = midtape STape (table@〈grid,false〉::〈curc,false〉::curconfig@〈grid,false〉::ls)
- 〈grid,false〉 rs →
- legal_tape ls 〈curc,false〉 rs →
- ∀t1'.t1' = lift_tape ls 〈curc,false〉 rs →
- ∃ls1,rs1,newc.
- (t2 = midtape STape ls1 〈grid,false〉 (reverse ? curconfig@〈newc,false〉::
- 〈grid,false〉::reverse ? table@〈grid,false〉::rs1) ∧
- lift_tape ls1 〈newc,false〉 rs1 =
- tape_move_left STape ls 〈curc,false〉 rs ∧ legal_tape ls1 〈newc,false〉 rs1).
-
-lemma mtl_concrete_to_abstract :
- ∀t1,t2.R_move_tape_l t1 t2 → R_move_tape_l_abstract t1 t2.
-#t1 #t2 whd in ⊢(%→?); #Hconcrete
-#rs #n #table #curc #curconfig #ls #Hcurc #Hcurconfig #Htable #Ht1
-* * * #Hnomarks #Hbits #Hcurc #Hlegal #t1' #Ht1'
-cases (Hconcrete … Htable ? Ht1) //
-[ * #Hls #Ht2
- @(ex_intro ?? [])
- @(ex_intro ?? (〈curc,false〉::rs))
- @(ex_intro ?? null) %
- [ %
- [ >Ht2 %
- | >Hls % ]
- | % [ % [ %
- [ #x #Hx cases (orb_true_l … Hx) #Hx'
- [ >(\P Hx') %
- | @Hnomarks >Hls @Hx' ]
- | #x #Hx cases (orb_true_l … Hx) #Hx'
- [ >(\P Hx') //
- | @Hbits >Hls @Hx' ]]
- | % ]
- | % %2 % ]
- ]
-| * * #l0 #bl0 * #ls0 * #Hls
- cut (bl0 = false)
- [ @(Hnomarks 〈l0,bl0〉) @memb_cons @memb_append_l1 >Hls @memb_hd]
- #Hbl0 >Hbl0 in Hls; #Hls #Ht2
- @(ex_intro ?? ls0) @(ex_intro ?? (〈curc,false〉::rs))
- @(ex_intro ?? l0) %
- [ %
- [ >Ht2 %
- | >Hls >lift_tape_not_null
- [ %
- | @bit_not_null @(Hbits 〈l0,false〉) >Hls @memb_append_l1 @memb_hd ] ]
- | % [ % [ %
- [ #x #Hx cases (orb_true_l … Hx) #Hx'
- [ >(\P Hx') %
- | cases (memb_append … Hx') #Hx'' @Hnomarks
- [ >Hls @memb_cons @memb_cons @(memb_append_l1 … Hx'')
- | cases (orb_true_l … Hx'') #Hx'''
- [ >(\P Hx''') @memb_hd
- | @memb_cons @(memb_append_l2 … Hx''')]
- ]
- ]
- | whd in ⊢ (?%); #x #Hx cases (memb_append … Hx) #Hx'
- [ @Hbits >Hls @memb_cons @(memb_append_l1 … Hx')
- | cases (orb_true_l … Hx') #Hx''
- [ >(\P Hx'') //
- | @Hbits @(memb_append_l2 … Hx'')
- ]]]
- | whd in ⊢ (??%?); >(Hbits 〈l0,false〉) //
- @memb_append_l1 >Hls @memb_hd ]
- | % % % #Hl0 lapply (Hbits 〈l0,false〉?)
- [ @memb_append_l1 >Hls @memb_hd
- | >Hl0 normalize #Hfalse destruct (Hfalse)
- ] ] ]
-| #x #Hx @Hbits @memb_append_l1 @Hx ]
-qed.
-
-lemma sem_move_tape_l_abstract : Realize … move_tape_l R_move_tape_l_abstract.
-@(Realize_to_Realize … mtl_concrete_to_abstract) //
-qed.
-
-lemma sem_move_tape_r_abstract : Realize … move_tape_r R_move_tape_r_abstract.
-@(Realize_to_Realize … mtr_concrete_to_abstract) //
-qed.
-
-(*
- t1 = ls # cs c # table # rs
-
- let simt ≝ lift_tape ls c rs in
- let simt' ≝ move_left simt' in
-
- t2 = left simt'# cs (sim_current_of_tape simt') # table # right simt'
-*)
-
-(*
-definition R_move
-
-definition R_exec_move ≝ λt1,t2.
- ∀ls,current,table1,newcurrent,table2,rs.
- t1 = midtape STape (current@〈grid,false〉::ls) 〈grid,false〉
- (table1@〈comma,true〉::newcurrent@〈comma,false〉::move::table2@
- 〈grid,false〉::rs) →
- table_TM (table1@〈comma,false〉::newcurrent@〈comma,false〉::move::table2) →
- t2 = midtape
-*)
-
-(*
-
-step :
-
-if is_true(current) (* current state is final *)
- then nop
- else
- init_match;
- match_tuple;
- if is_marked(current) = false (* match ok *)
- then exec_move;
- else sink;
-
-*)
-
-
-definition move_tape ≝
- ifTM ? (test_char ? (λc:STape.c == 〈bit false,false〉))
- (* spostamento a sinistra: verificare se per caso non conviene spostarsi
- sulla prima grid invece dell'ultima *)
- (adv_to_mark_r ? (λc:STape.is_grid (\fst c)) · move_tape_l)
- (ifTM ? (test_char ? (λc:STape.c == 〈bit true,false〉))
- (adv_to_mark_r ? (λc:STape.is_grid (\fst c)) · move_tape_r)
- (adv_to_mark_l ? (λc:STape.is_grid (\fst c)) ·
- move_l … · adv_to_mark_l ? (λc:STape.is_grid (\fst c)))
- tc_true) tc_true.
-
-definition R_move_tape ≝ λt1,t2.
- ∀rs,n,table1,mv,table2,curc,curconfig,ls.
- bit_or_null mv = true → only_bits_or_nulls curconfig →
- (is_bit mv = true → is_bit curc = true) →
- table_TM n (reverse ? table1@〈mv,false〉::table2) →
- t1 = midtape STape (table1@〈grid,false〉::〈curc,false〉::curconfig@〈grid,false〉::ls)
- 〈mv,false〉 (table2@〈grid,false〉::rs) →
- legal_tape ls 〈curc,false〉 rs →
- ∀t1'.t1' = lift_tape ls 〈curc,false〉 rs →
- ∃ls1,rs1,newc.
- legal_tape ls1 〈newc,false〉 rs1 ∧
- (t2 = midtape STape ls1 〈grid,false〉 (reverse ? curconfig@〈newc,false〉::
- 〈grid,false〉::reverse ? table1@〈mv,false〉::table2@〈grid,false〉::rs1) ∧
- ((mv = bit false ∧ lift_tape ls1 〈newc,false〉 rs1 = tape_move_left STape ls 〈curc,false〉 rs) ∨
- (mv = bit true ∧ lift_tape ls1 〈newc,false〉 rs1 = tape_move_right STape ls 〈curc,false〉 rs) ∨
- (mv = null ∧ ls1 = ls ∧ rs1 = rs ∧ curc = newc))).
-
-lemma sem_move_tape : Realize ? move_tape R_move_tape.
-#intape
-cases (sem_if ? (test_char ??) … tc_true (sem_test_char ? (λc:STape.c == 〈bit false,false〉))
- (sem_seq … (sem_adv_to_mark_r ? (λc:STape.is_grid (\fst c))) sem_move_tape_l_abstract)
- (sem_if ? (test_char ??) … tc_true (sem_test_char ? (λc:STape.c == 〈bit true,false〉))
- (sem_seq … (sem_adv_to_mark_r ? (λc:STape.is_grid (\fst c))) sem_move_tape_r_abstract)
- (sem_seq … (sem_adv_to_mark_l ? (λc:STape.is_grid (\fst c)))
- (sem_seq … (sem_move_l …) (sem_adv_to_mark_l ? (λc:STape.is_grid (\fst c)))))) intape)
-#k * #outc * #Hloop #HR
-@(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#rs #n #table1 #mv #table2 #curc #curconfig #ls
-#Hmv #Hcurconfig #Hmvcurc #Htable #Hintape #Htape #t1' #Ht1'
-generalize in match HR; -HR *
-[* #ta * whd in ⊢ (%→?); * * #c * >Hintape normalize in ⊢ (%→?);
- #Hdes destruct (Hdes) #eqmv
- cut (mv = bit false) [lapply (\P eqmv) #Hdes destruct (Hdes) %] -eqmv #eqmv #Hta
- * #tb * whd in ⊢ (%→?); #Htb cases (proj2 ?? Htb … Hta) -Htb -Hta
- [ * >eqmv normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
- * * #_ #Htb #_ lapply (Htb ??? (refl ??) (refl …) ?)
- [ @daemon ] -Htb >append_cons <associative_append #Htb
- whd in ⊢ (%→?); #Houtc lapply (Houtc … Htb … Ht1') //
- [ >reverse_append >reverse_append >reverse_reverse @Htable |@Hmvcurc >eqmv % ||]
- -Houtc -Htb * #ls1 * #rs1 * #newc * * #Houtc #Hnewtape #Hnewtapelegal
- @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) %
- [ //
- | %
- [ >Houtc >reverse_append >reverse_append >reverse_reverse
- >associative_append >associative_append %
- | % % % // ]
- ]
-| * #ta * whd in ⊢ (%→?); * >Hintape #Hcneq #Hta
- cut (mv ≠ bit false)
- [lapply (\Pf (Hcneq … (refl …))) @not_to_not #Heq >Heq % ] -Hcneq #Hcneq
- *
- [* #tb * * * #ca >Hta -Hta * normalize in ⊢ (%→?); #Hdes destruct (Hdes) #eqmv
- cut (mv = bit true) [lapply (\P eqmv) #Hdes destruct (Hdes) %] -eqmv #eqmv
- #Htb * #tc * whd in ⊢ (%→?); #Htc cases (proj2 ?? Htc … Htb) -Htc
- [ * >(bit_or_null_not_grid … Hmv) #Hfalse destruct (Hfalse) ]
- * * #_ #Htc lapply (Htc ???(refl ??) (refl ??) ?)
- [ @daemon ] -Htc >append_cons <associative_append #Htc
- #_ whd in ⊢ (%→?); #Houtc lapply (Houtc ? n … Htc … Ht1') //
- [ >reverse_append >reverse_append >reverse_reverse @Htable
- | @Hmvcurc @daemon]
- -Houtc -Htc * #ls1 * #rs1 * #newc * * #Houtc #Hnewtape #Hnewtapelegal
- @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) %
- [ //
- | %
- [ >Houtc >reverse_append >reverse_append >reverse_reverse
- >associative_append >associative_append %
- | % %2 % //
- ]
- ]
- |* #tb * whd in ⊢ (%→?); * >Hta #Hcneq' #Htb
- cut (mv ≠ bit true)
- [lapply (\Pf (Hcneq' … (refl …))) @not_to_not #Heq >Heq % ] -Hcneq' #Hcneq'
- * #tc * whd in ⊢ (%→?); #Htc cases (proj2 ?? Htc … Htb) -Htc
- #_ #Htc cases (Htc (bit_or_null_not_grid … Hmv)) -Htc #Htc #_
- lapply (Htc … (refl ??) (refl ??) ?) [@daemon] -Htc #Htc
- * #td * whd in ⊢ (%→?); #Htd lapply (proj2 ?? Htd … Htc) -Htd -Htc
- whd in ⊢ (???%→?); #Htd whd in ⊢ (%→?); #Houtc cases (proj2 ?? Houtc … Htd)
- -Houtc #_ #Houtc cases (Houtc ?)
- [2: cases Htape * * #_ #_ #Hcurc #_ >(bit_or_null_not_grid … Hcurc) %]
- #Houtc #_ lapply (Houtc … (refl ??) (refl ??) ?) [@daemon] -Houtc #Houtc
- @(ex_intro ?? ls) @(ex_intro ?? rs) @(ex_intro ?? curc) %
- [ //
- | %
- [ @Houtc
- | %2 % // % // % //
- generalize in match Hcneq; generalize in match Hcneq';
- cases mv in Hmv; normalize //
- [ * #_ normalize [ #Hfalse @False_ind cases Hfalse /2/ | #_ #Hfalse @False_ind cases Hfalse /2/ ]
- |*: #Hfalse destruct (Hfalse) ]
- ]
- ]
- ]
-]
-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/universal/alphabet.ma".
-include "turing/mono.ma".
-
-(************************* turning DeqMove into a DeqSet **********************)
-definition move_eq ≝ λm1,m2:move.
- match m1 with
- [R ⇒ match m2 with [R ⇒ true | _ ⇒ false]
- |L ⇒ match m2 with [L ⇒ true | _ ⇒ false]
- |N ⇒ match m2 with [N ⇒ true | _ ⇒ false]].
-
-lemma move_eq_true:∀m1,m2.
- move_eq m1 m2 = true ↔ m1 = m2.
-*
- [* normalize [% #_ % |2,3: % #H destruct ]
- |* normalize [1,3: % #H destruct |% #_ % ]
- |* normalize [1,2: % #H destruct |% #_ % ]
-qed.
-
-definition DeqMove ≝ mk_DeqSet move move_eq move_eq_true.
-
-unification hint 0 ≔ ;
- X ≟ DeqMove
-(* ---------------------------------------- *) ⊢
- move ≡ carr X.
-
-unification hint 0 ≔ m1,m2;
- X ≟ DeqMove
-(* ---------------------------------------- *) ⊢
- move_eq m1 m2 ≡ eqb X m1 m2.
-
-
-(************************ turning DeqMove into a FinSet ***********************)
-definition move_enum ≝ [L;R;N].
-
-lemma move_enum_unique: uniqueb ? [L;R;N] = true.
-// qed.
-
-lemma move_enum_complete: ∀x:move. memb ? x [L;R;N] = true.
-* // qed.
-
-definition FinMove ≝
- mk_FinSet DeqMove [L;R;N] move_enum_unique move_enum_complete.
-
-unification hint 0 ≔ ;
- X ≟ FinMove
-(* ---------------------------------------- *) ⊢
- move ≡ FinSetcarr X.
-
-(*************************** normal Turing Machines ***************************)
-
-(* A normal turing machine is just an ordinary machine where:
- 1. the tape alphabet is bool
- 2. the finite state are supposed to be an initial segment of the natural
- numbers.
- Formally, it is specified by a record with the number n of states, a proof
- that n is positive, the transition function and the halting function.
-*)
-
-definition trans_source ≝ λn.FinProd (initN n) (FinOption FinBool).
-definition trans_target ≝ λn.FinProd (initN n) (FinOption (FinProd FinBool FinMove)).
-
-record normalTM : Type[0] ≝
-{ no_states : nat;
- pos_no_states : (0 < no_states);
- ntrans : trans_source no_states → trans_target no_states;
- nhalt : initN no_states → bool
-}.
-
-(* A normal machine is just a special case of Turing Machine. *)
-
-definition normalTM_to_TM ≝ λM:normalTM.
- mk_TM FinBool (initN (no_states M))
- (ntrans M) (mk_Sig ?? 0 (pos_no_states M)) (nhalt M).
-
-coercion normalTM_to_TM.
-
-definition nconfig ≝ λn. config FinBool (initN n).
-
-(******************************** tuples **************************************)
-
-(* By general results on FinSets we know that every function f between two
-finite sets A and B can be described by means of a finite graph of pairs
-〈a,f a〉. Hence, the transition function of a normal turing machine can be
-described by a finite set of tuples 〈i,c〉,〈j,action〉〉 of the following type:
- (Nat_to n × (option FinBool)) × (Nat_to n × (option (FinBool × move))).
-Unfortunately this description is not suitable for a Universal Machine, since
-such a machine must work with a fixed alphabet, while the size on n is unknown.
-Hence, we must pass from natural numbers to a representation for them on a
-finitary, e.g. binary, alphabet. In general, we shall associate
-to a pair 〈〈i,c〉,〈j,action〉〉 a tuple with the following syntactical structure
- |w_ix,w_jy,z
-where
-1. "|" and "," are special characters used as delimiters;
-2. w_i and w_j are list of booleans representing the states $i$ and $j$;
-3. x is special symbol null if c=None and is a if c=Some a
-4. y and z are both null if action = None, and are equal to b,m' if
- action = Some b,m;
-5. finally, m' = 0 if m = L, m' = 1 if m=R and m' = null if m = N
-
-As a minor, additional complication, we shall suppose that every characters is
-decorated by an additonal bit, normally set to false, to be used as a marker.
-*)
-
-definition mk_tuple ≝ λqin,cin,qout,cout,mv.
- 〈bar,false〉::qin@cin::〈comma,false〉::qout@cout::〈comma,false〉::[mv].
-
-(* by definition, a tuple is not marked *)
-definition tuple_TM : nat → list STape → Prop ≝
- λn,t.∃qin,cin,qout,cout,mv.
- no_marks qin ∧ no_marks qout ∧
- only_bits qin ∧ only_bits qout ∧
- bit_or_null cin = true ∧ bit_or_null cout = true ∧ bit_or_null mv = true ∧
- (cout = null → mv = null) ∧
- |qin| = n ∧ |qout| = n ∧
- t = mk_tuple qin 〈cin,false〉 qout 〈cout,false〉 〈mv,false〉.
-
-(***************************** state encoding *********************************)
-(* p < n is represented with a list of bits of lenght n with the p-th bit from
-left set to 1. An additional intial bit is set to 1 if the state is final and
-to 0 otherwise. *)
-
-let rec to_bitlist n p: list bool ≝
- match n with [ O ⇒ [ ] | S q ⇒ (eqb p q)::to_bitlist q p].
-
-let rec from_bitlist l ≝
- match l with
- [ nil ⇒ 0 (* assert false *)
- | cons b tl ⇒ if b then |tl| else from_bitlist tl].
-
-lemma bitlist_length: ∀n,p.|to_bitlist n p| = n.
-#n elim n normalize //
-qed.
-
-lemma bitlist_inv1: ∀n,p.p<n → from_bitlist (to_bitlist n p) = p.
-#n elim n normalize -n
- [#p #abs @False_ind /2/
- |#n #Hind #p #lepn
- cases (le_to_or_lt_eq … (le_S_S_to_le … lepn))
- [#ltpn lapply (lt_to_not_eq … ltpn) #Hpn
- >(not_eq_to_eqb_false … Hpn) normalize @Hind @ltpn
- |#Heq >(eq_to_eqb_true … Heq) normalize <Heq //
- ]
- ]
-qed.
-
-lemma bitlist_lt: ∀l. 0 < |l| → from_bitlist l < |l|.
-#l elim l normalize // #b #tl #Hind cases b normalize //
-#Htl cases (le_to_or_lt_eq … (le_S_S_to_le … Htl)) -Htl #Htl
- [@le_S_S @lt_to_le @Hind //
- |cut (tl=[ ]) [/2 by append_l2_injective/] #eqtl >eqtl @le_n
- ]
-qed.
-
-definition nat_of: ∀n. Nat_to n → nat.
-#n normalize * #p #_ @p
-qed.
-
-definition bits_of_state ≝ λn.λh:Nat_to n → bool.λs:Nat_to n.
- h s::(to_bitlist n (nat_of n s)).
-
-definition m_bits_of_state ≝ λn.λh.λp.
- map ? (unialpha×bool) (λx.〈bit x,false〉) (bits_of_state n h p).
-
-lemma no_marks_bits_of_state : ∀n,h,p. no_marks (m_bits_of_state n h p).
-#n #h #p #x whd in match (m_bits_of_state n h p);
-#H cases (orb_true_l … H) -H
- [#H >(\P H) %
- |elim (to_bitlist n (nat_of n p))
- [whd in ⊢ ((??%?)→?); #H destruct
- |#b #l #Hind #H cases (orb_true_l … H) -H #H
- [>(\P H) % |@Hind @H]
- ]
- ]
-qed.
-
-lemma only_bits_bits_of_state : ∀n,h,p. only_bits (m_bits_of_state n h p).
-#n #h #p #x whd in match (m_bits_of_state n h p);
-#H cases (orb_true_l … H) -H
- [#H >(\P H) %
- |elim (to_bitlist n (nat_of n p))
- [whd in ⊢ ((??%?)→?); #H destruct
- |#b #l #Hind #H cases (orb_true_l … H) -H #H
- [>(\P H) % |@Hind @H ]
- ]
- ]
-qed.
-
-(******************************** action encoding *****************************)
-definition low_action ≝ λaction.
- match action with
- [ None ⇒ 〈null,null〉
- | Some act ⇒ let 〈na,m〉 ≝ act in
- match m with
- [ R ⇒ 〈bit na,bit true〉
- | L ⇒ 〈bit na,bit false〉
- | N ⇒ 〈bit na,null〉]
- ].
-
-(******************************** tuple encoding ******************************)
-definition tuple_type ≝ λn.
- (Nat_to n × (option FinBool)) × (Nat_to n × (option (FinBool × move))).
-
-definition tuple_encoding ≝ λn.λh:Nat_to n→bool.
- λp:tuple_type n.
- let 〈inp,outp〉 ≝ p in
- let 〈q,a〉 ≝ inp in
- let cin ≝ match a with [ None ⇒ null | Some b ⇒ bit b ] in
- let 〈qn,action〉 ≝ outp in
- let 〈cout,mv〉 ≝ low_action action in
- let qin ≝ m_bits_of_state n h q in
- let qout ≝ m_bits_of_state n h qn in
- mk_tuple qin 〈cin,false〉 qout 〈cout,false〉 〈mv,false〉.
-
-(*
-definition WFTuple_conditions ≝
- λn,qin,cin,qout,cout,mv.
- no_marks qin ∧ no_marks qout ∧ (* queste fuori ? *)
- only_bits qin ∧ only_bits qout ∧
- bit_or_null cin = true ∧ bit_or_null cout = true ∧ bit_or_null mv = true ∧
- (cout = null → mv = null) ∧
- |qin| = n ∧ |qout| = n. *)
-
-lemma is_tuple: ∀n,h,p. tuple_TM (S n) (tuple_encoding n h p).
-#n #h * * #q #a * #qn #action
-@(ex_intro … (m_bits_of_state n h q))
-letin cin ≝ match a with [ None ⇒ null | Some b ⇒ bit b ]
-@(ex_intro … cin)
-@(ex_intro … (m_bits_of_state n h qn))
-letin cout ≝
- match action with
- [ None ⇒ null | Some act ⇒ bit (\fst act)]
-@(ex_intro … cout)
-letin mv ≝ match action with
- [ None ⇒ null
- | Some act ⇒
- match \snd act with
- [ R ⇒ bit true | L ⇒ bit false | N ⇒ null]
- ]
-@(ex_intro … mv)
-%[%[%[%[%[%[%[% /3/
- |whd in match cin ; cases a //]
- |whd in match cout; cases action //]
- |whd in match mv; cases action // * #b #m cases m //]
- |whd in match cout; whd in match mv; cases action
- [// | #act whd in ⊢ ((??%?)→?); #Hfalse destruct ]]
- |>length_map normalize @eq_f //]
- |>length_map normalize @eq_f //]
- |normalize cases a cases action normalize //
- [* #c #m cases m % |* #c #m #c1 cases m %]
- ]
-qed.
-
-definition tuple_length ≝ λn.2*n+6.
-
-lemma length_of_tuple: ∀n,t. tuple_TM n t →
- |t| = tuple_length n.
-#n #t * #qin * #cin * #qout * #cout * #mv *** #_ #Hqin #Hqout #eqt >eqt
-whd in match (mk_tuple ?????); normalize >length_append >Hqin -Hqin normalize
->length_append normalize >Hqout -Hqout //
-qed.
-
-definition tuples_list ≝ λn.λh.map … (λp.tuple_encoding n h p).
-
-(******************* general properties of encoding of tuples *****************)
-
-lemma no_grids_in_tuple : ∀n,l.tuple_TM n l → no_grids l.
-#n #l * #qin * #cin * #qout * #cout * #mv * * * * * * * * * *
-#_ #_ #Hqin #Hqout #Hcin #Hcout #Hmv #_ #_ #_ #Hl >Hl
-#c #Hc cases (orb_true_l … Hc) -Hc #Hc
-[ >(\P Hc) %
-| cases (memb_append … Hc) -Hc #Hc
-[ @bit_not_grid @(Hqin … Hc)
-| cases (orb_true_l … Hc) -Hc #Hc
-[ change with (c == 〈cin,false〉 = true) in Hc; >(\P Hc) @bit_or_null_not_grid //
-| cases (orb_true_l … Hc) -Hc #Hc
-[ change with (c == 〈comma,false〉 = true) in Hc; >(\P Hc) %
-| cases (memb_append …Hc) -Hc #Hc
-[ @bit_not_grid @(Hqout … Hc)
-| cases (orb_true_l … Hc) -Hc #Hc
-[ change with (c == 〈cout,false〉 = true) in Hc; >(\P Hc) @bit_or_null_not_grid //
-| cases (orb_true_l … Hc) -Hc #Hc
-[ change with (c == 〈comma,false〉 = true) in Hc; >(\P Hc) %
-| >(memb_single … Hc) @bit_or_null_not_grid @Hmv
-]]]]]]
-qed.
-
-lemma no_marks_in_tuple : ∀n,l.tuple_TM n l → no_marks l.
-#n #l * #qin * #cin * #qout * #cout * #mv * * * * * * * * * *
-#Hqin #Hqout #_ #_ #_ #_ #_ #_ #_ #_ #Hl >Hl
-#c #Hc cases (orb_true_l … Hc) -Hc #Hc
-[ >(\P Hc) %
-| cases (memb_append … Hc) -Hc #Hc
-[ @(Hqin … Hc)
-| cases (orb_true_l … Hc) -Hc #Hc
-[ change with (c == 〈cin,false〉 = true) in Hc; >(\P Hc) %
-| cases (orb_true_l … Hc) -Hc #Hc
-[ change with (c == 〈comma,false〉 = true) in Hc; >(\P Hc) %
-| cases (memb_append … Hc) -Hc #Hc
-[ @(Hqout … Hc)
-| cases (orb_true_l … Hc) -Hc #Hc
-[ change with (c == 〈cout,false〉 = true) in Hc; >(\P Hc) %
-| cases (orb_true_l … Hc) -Hc #Hc
-[ change with (c == 〈comma,false〉 = true) in Hc; >(\P Hc) %
-| >(memb_single … Hc) %
-]]]]]]
-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_____________________________________________________________*)
-
-
-(****************************** table of tuples *******************************)
-include "turing/universal/normalTM.ma".
-
-(* a well formed table is a list of tuples *)
-
-inductive table_TM (n:nat) : list STape → Prop ≝
-| ttm_nil : table_TM n []
-| ttm_cons : ∀t1,T.tuple_TM n t1 → table_TM n T → table_TM n (t1@T).
-
-lemma wftable: ∀n,h,l.table_TM (S n) (flatten ? (tuples_list n h l)).
-#n #h #l elim l // -l #a #tl #Hind
-whd in match (flatten … (tuples_list …));
-@ttm_cons //
-qed.
-
-(*********************** general properties of tables *************************)
-lemma no_grids_in_table: ∀n.∀l.table_TM n l → no_grids l.
-#n #l #t elim t
- [normalize #c #H destruct
- |#t1 #t2 #Ht1 #Ht2 #IH lapply (no_grids_in_tuple … Ht1) -Ht1 #Ht1 #x #Hx
- cases (memb_append … Hx) -Hx #Hx
- [ @(Ht1 … Hx)
- | @(IH … Hx) ] ]
-qed.
-
-lemma no_marks_in_table: ∀n.∀l.table_TM n l → no_marks l.
-#n #l #t elim t
- [normalize #c #H destruct
- |#t1 #t2 #Ht1 #Ht2 #IH lapply (no_marks_in_tuple … Ht1) -Ht1 #Ht1 #x #Hx
- cases (memb_append … Hx) -Hx #Hx
- [ @(Ht1 … Hx)
- | @(IH … Hx) ] ]
-qed.
-
-axiom last_of_table: ∀n,l,b.¬ table_TM n (l@[〈bar,b〉]).
-
-(************************** matching in a table *******************************)
-inductive match_in_table (n:nat) (qin:list STape) (cin: STape)
- (qout:list STape) (cout:STape) (mv:STape)
-: list STape → Prop ≝
-| mit_hd :
- ∀tb.
- tuple_TM n (mk_tuple qin cin qout cout mv) →
- match_in_table n qin cin qout cout mv
- (mk_tuple qin cin qout cout mv @tb)
-| mit_tl :
- ∀qin0,cin0,qout0,cout0,mv0,tb.
- tuple_TM n (mk_tuple qin0 cin0 qout0 cout0 mv0) →
- match_in_table n qin cin qout cout mv tb →
- match_in_table n qin cin qout cout mv
- (mk_tuple qin0 cin0 qout0 cout0 mv0@tb).
-
-lemma tuple_to_match: ∀n,h,l,qin,cin,qout,cout,mv,p.
- p = mk_tuple qin cin qout cout mv
- → mem ? p (tuples_list n h l) →
- match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_list n h l)).
-#n #h #l #qin #cin #qout #cout #mv #p
-#Hp elim l
- [whd in ⊢ (%→?); @False_ind
- |#p1 #tl #Hind *
- [#H whd in match (tuples_list ???);
- <H >Hp @mit_hd //
- |#H whd in match (tuples_list ???);
- cases (is_tuple n h p1) #qin1 * #cin1 * #qout1 * #cout1 * #mv1
- * #_ #Htuplep1 >Htuplep1 @mit_tl // @Hind //
- ]
- ]
-qed.
-
-axiom match_decomp: ∀n,l,qin,cin,qout,cout,mv.
- match_in_table (S n) qin cin qout cout mv l →
- ∃l1,l2. l = l1@(mk_tuple qin cin qout cout mv)@l2 ∧
- (∃q.|l1| = (tuple_length (S n))*q) ∧
- tuple_TM (S n) (mk_tuple qin cin qout cout mv).
-
-lemma match_to_tuples_list: ∀n,h,l,qin,cin,qout,cout,mv.
- match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_list n h l)) →
- ∃p. p = mk_tuple qin cin qout cout mv ∧ mem ? p (tuples_list n h l).
-#n #h #l #qin #cin #qout #cout #mv #Hmatch
-@(ex_intro … (mk_tuple qin cin qout cout mv)) % //
-cases (match_decomp … Hmatch) #l1 * #l2 * * #Hflat #Hlen #Htuple
-@(flatten_to_mem … Hflat … Hlen)
- [//
- |#x #memx @length_of_tuple
- cases (mem_map ????? memx) #t * #memt #Ht <Ht //
- |@(length_of_tuple … Htuple)
- ]
-qed.
-
-lemma match_to_tuple: ∀n,h,l,qin,cin,qout,cout,mv.
- match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_list n h l)) →
- ∃p. tuple_encoding n h p = mk_tuple qin cin qout cout mv ∧ mem ? p l.
-#n #h #l #qin #cin #qout #cout #mv #Hmatch
-cases (match_to_tuples_list … Hmatch)
-#p * #eqp #memb
-cases(mem_map … (λp.tuple_encoding n h p) … memb)
-#p1 * #Hmem #H @(ex_intro … p1) % /2/
-qed.
-
-lemma match_to_trans:
- ∀n.∀trans: trans_source n → trans_target n.
- ∀h,qin,cin,qout,cout,mv.
- match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_list n h (graph_enum ?? trans))) →
- ∃s,t. tuple_encoding n h 〈s,t〉 = mk_tuple qin cin qout cout mv
- ∧ trans s = t.
-#n #trans #h #qin #cin #qout #cout #mv #Hmatch
-cases (match_to_tuple … Hmatch) -Hmatch * #s #t * #Heq #Hmem
-@(ex_intro … s) @(ex_intro … t) % // @graph_enum_correct
-@mem_to_memb @Hmem
-qed.
-
-lemma trans_to_match:
- ∀n.∀h.∀trans: trans_source n → trans_target n.
- ∀s,t,qin,cin,qout,cout,mv. trans s = t →
- tuple_encoding n h 〈s,t〉 = mk_tuple qin cin qout cout mv →
- match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_list n h (graph_enum ?? trans))).
-#n #h #trans #inp #outp #qin #cin #qout #cout #mv #Htrans #Htuple
-@(tuple_to_match … (refl…)) <Htuple @mem_map_forward
-@(memb_to_mem (FinProd (trans_source n) (trans_target n)))
-@graph_enum_complete //
-qed.
-
-(*
-axiom append_eq_tech1 :
- ∀A,l1,l2,l3,l4.l1@l2 = l3@l4 → |l1| < |l3| → ∃la:list A.l1@la = l3.
-axiom append_eq_tech2 :
- ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → memb A a l4 = false → ∃la:list A.l3 = l1@a::la.
-axiom list_decompose_cases :
- ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → ∃la,lb:list A.l3 = la@a::lb ∨ l4 = la@a::lb.
-axiom list_decompose_l :
- ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → memb A a l4 = false →
- ∃la,lb.l2 = la@lb ∧ l3 = l1@a::la.*)
-
-lemma list_decompose_r :
- ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → memb A a l3 = false →
- ∃la,lb.l1 = la@lb ∧ l4 = lb@a::l2.
-#A #l1 #l2 #l3 generalize in match l1; generalize in match l2; elim l3
- [normalize #l1 #l2 #l4 #a #H #_ @(ex_intro … []) @(ex_intro … l2) /2/
- |#b #tl #Hind #l1 #l2 #l4 #a cases l2
- [normalize #Heq destruct >(\b (refl … b)) normalize #Hfalse destruct
- |#c #tl2 whd in ⊢ ((??%%)→?); #Heq destruct #Hmema
- cases (Hind l1 tl2 l4 a ??)
- [#l5 * #l6 * #eql #eql4
- @(ex_intro … (b::l5)) @(ex_intro … l6) % /2/
- |@e0
- |cases (true_or_false (memb ? a tl)) [2://]
- #H @False_ind @(absurd ?? not_eq_true_false)
- <Hmema @sym_eq @memb_cons //
- ]
- ]
- ]
-qed.
-
-(*axiom list_decompose_memb :
- ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → |l1| < |l3| → memb A a l3 = true.*)
-
-lemma table_invert_r : ∀n,t,T.
- tuple_TM n t → table_TM n (t@T) → table_TM n T.
-#n #t #T #Htuple #Htable inversion Htable
-[ cases Htuple #qin * #cin * #qout * #cout * #mv * #_ #Ht >Ht
- normalize #Hfalse destruct (Hfalse)
-| #t0 #T0 #Htuple0 #Htable0 #_ #Heq
- lapply (append_l2_injective ?????? Heq)
- [ >(length_of_tuple … Htuple) >(length_of_tuple … Htuple0) % ]
- -Heq #Heq destruct (Heq) // ]
-qed.
-
-lemma match_in_table_to_tuple :
- ∀n,T,qin,cin,qout,cout,mv.
- match_in_table n qin cin qout cout mv T → table_TM n T →
- tuple_TM n (mk_tuple qin cin qout cout mv).
-#n #T #qin #cin #qout #cout #mv #Hmatch elim Hmatch
-[ //
-| #qin0 #cin0 #qout0 #cout0 #mv0 #tb #Htuple #Hmatch #IH #Htable
- @IH @(table_invert_r ???? Htable) @Htuple
-]
-qed.
-
-lemma match_in_table_append :
- ∀n,T,qin,cin,qout,cout,mv,t.
- tuple_TM n t →
- match_in_table n qin cin qout cout mv (t@T) →
- t = mk_tuple qin cin qout cout mv ∨ match_in_table n qin cin qout cout mv T.
-#n #T #qin #cin #qout #cout #mv #t #Ht #Hmatch inversion Hmatch
-[ #T0 #H #H1 % >(append_l1_injective … H1) //
- >(length_of_tuple … Ht) >(length_of_tuple … H) %
-| #qin0 #cin0 #qout0 #cout0 #mv0 #T0 #H #H1 #_ #H2 %2
- >(append_l2_injective … H2) // >(length_of_tuple … Ht) >(length_of_tuple … H) %
-]
-qed.
-
-lemma generic_match_to_match_in_table_tech :
- ∀n,t,T0,T1,T2.tuple_TM n t → table_TM n (T1@〈bar,false〉::T2) →
- t@T0 = T1@〈bar,false〉::T2 → T1 = [] ∨ ∃T3.T1 = t@T3.
-#n #t #T0 #T1 #T2 #Ht cases T1
-[ #_ #_ % %
-| normalize #c #T1c #Htable #Heq %2
- cases Ht in Heq; #qin * #cin * #qout * #cout * #mv **********
- #Hqin1 #Hqout1 #Hqin2 #Hqout2 #Hcin #Hcout #Hmv #Hcoutmv #Hqinlen #Hqoutlen
- #Heqt >Heqt whd in ⊢ (??%%→?); #Ht lapply (cons_injective_r ????? Ht)
- #Ht' cases (list_decompose_r STape … (sym_eq … Ht') ?)
- [ #la * #lb * #HT1c #HT0 %{lb} >HT1c @(eq_f2 ??? (append ?) (c::la)) //
- >HT0 in Ht'; >HT1c >associative_append in ⊢ (???%→?); #Ht'
- <(append_l1_injective_r … Ht') // <(cons_injective_l ????? Ht) %
- |@(noteq_to_eqnot ? true) @(not_to_not … not_eq_true_false) #Hbar @sym_eq
- cases (memb_append … Hbar) -Hbar #Hbar
- [@(Hqin2 … Hbar)
- |cases (orb_true_l … Hbar) -Hbar
- [#Hbar lapply (\P Hbar) -Hbar #Hbar destruct (Hbar) @Hcin
- |whd in ⊢ ((??%?)→?); #Hbar cases (memb_append … Hbar) -Hbar #Hbar
- [@(Hqout2 … Hbar)
- |cases (orb_true_l … Hbar) -Hbar
- [#Hbar lapply (\P Hbar) -Hbar #Hbar destruct (Hbar) @Hcout
- |#Hbar cases (orb_true_l … Hbar) -Hbar
- [whd in ⊢ ((??%?)→?); #Hbar @Hbar
- |#Hbar lapply (memb_single … Hbar) -Hbar #Hbar destruct (Hbar) @Hmv
- ]
- ]
- ]
- ]
- ]
- ]
-qed.
-
-lemma generic_match_to_match_in_table :
- ∀n,T.table_TM n T →
- ∀qin,cin,qout,cout,mv.|qin| = n → |qout| = n →
- only_bits qin → only_bits qout →
- bit_or_null (\fst cin) = true → bit_or_null (\fst cout) = true →
- bit_or_null (\fst mv) = true →
- ∀t1,t2.
- T = (t1@〈bar,false〉::qin@cin::〈comma,false〉::qout@cout::〈comma,false〉::[mv])@t2 →
- match_in_table n qin cin qout cout mv T.
-#n #T #Htable #qin #cin #qout #cout #mv #Hlenqin #Hlenqout
-#Hqinbits #Hqoutbits #Hcin #Hcout #Hmv
-elim Htable
-[ * [ #t2 normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
- | #c0 #t1 #t2 normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
-| #tuple #T0 #H1 #Htable0#IH #t1 #t2 #HT cases H1 #qin0 * #cin0 * #qout0 * #cout0 * #mv0
- * * * * * * * * * *
- #Hqin0marks #Hqout0marks #Hqin0bits #Hqout0bits #Hcin0 #Hcout0 #Hmv0 #Hcout0mv0
- #Hlenqin0 #Hlenqout0 #Htuple
- lapply (generic_match_to_match_in_table_tech n ? T0 t1
- (qin@cin::〈comma,false〉::qout@[cout;〈comma,false〉;mv]@t2) H1) #Htmp
- >Htuple in H1; #H1
- lapply (ttm_cons … T0 H1 Htable0) <Htuple in ⊢ (%→?); >HT
- >associative_append normalize >associative_append normalize
- >associative_append #Htable cases (Htmp Htable ?)
- [ #Ht1 >Htuple in HT; >Ht1 normalize in ⊢ (??%%→?);
- >associative_append >associative_append #HT
- cut (qin0 = qin ∧ (〈cin0,false〉 = cin ∧ (qout0 = qout ∧
- (〈cout0,false〉 = cout ∧ (〈mv0,false〉 = mv ∧ T0 = t2)))))
- [ lapply (cons_injective_r ????? HT) -HT #HT
- lapply (append_l1_injective … HT) [ >Hlenqin @Hlenqin0 ]
- #Hqin % [ @Hqin ] -Hqin
- lapply (append_l2_injective … HT) [ >Hlenqin @Hlenqin0 ] -HT #HT
- lapply (cons_injective_l ????? HT) #Hcin % [ @Hcin ] -Hcin
- lapply (cons_injective_r ????? HT) -HT #HT
- lapply (cons_injective_r ????? HT) -HT
- >associative_append >associative_append #HT
- lapply (append_l1_injective … HT) [ >Hlenqout @Hlenqout0 ]
- #Hqout % [ @Hqout ] -Hqout
- lapply (append_l2_injective … HT) [ >Hlenqout @Hlenqout0 ] -HT normalize #HT
- lapply (cons_injective_l ????? HT) #Hcout % [ @Hcout ] -Hcout
- lapply (cons_injective_r ????? HT) -HT #HT
- lapply (cons_injective_r ????? HT) -HT #HT
- lapply (cons_injective_l ????? HT) #Hmv % [ @Hmv ] -Hmv
- @(cons_injective_r ????? HT) ]
- -HT * #Hqin * #Hcin * #Hqout * #Hcout * #Hmv #HT0
- >(?:〈bar,false〉::qin0@(〈cin0,false〉::〈comma,false〉::qout0@
- [〈cout0,false〉;〈comma,false〉;〈mv0,false〉])@T0 = tuple@T0)
- [ >Htuple >Hqin >Hqout >Hcin >Hcout >Hmv % //
- | >Htuple normalize >associative_append normalize >associative_append
- normalize >associative_append % ]
- | * #T3 #HT3 >HT3 in HT; >associative_append; >associative_append #HT
- lapply (append_l2_injective … HT) // -HT #HT %2 //
- @(IH T3 t2) >HT >associative_append %
- |>HT >associative_append normalize >associative_append normalize
- >associative_append % ]
-]
-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_____________________________________________________________*)
-
-
-(* COMPARE BIT
-
-*)
-
-include "turing/universal/copy.ma".
-include "turing/universal/move_tape.ma".
-include "turing/universal/match_machines.ma".
-
-(*
-
-step :
-
-if is_true(current) (* current state is final *)
- then nop
- else
- (* init_match *)
- mark;
- adv_to_grid_r;
- move_r;
- mark;
- move_l;
- adv_to_mark_l
- (* /init_match *)
- match_tuple;
- if is_marked(current) = false (* match ok *)
- then
- (* init_copy *)
- move_l;
- init_current;
- move_r;
- adv_to_mark_r;
- adv_mark_r;
- (* /init_copy *)
- copy;
- move_r;
- (* move_tape *)
- by cases on current:
- case bit false: move_tape_l
- case bit true: move_tape_r
- case null: adv_to_grid_l; move_l; adv_to_grid_l;
- move_r;
- (* /move_tape *)
- else sink;
-
-*)
-
-definition init_match ≝
- mark ? · adv_to_mark_r ? (λc:STape.is_grid (\fst c)) · move_r ? ·
- move_r ? · mark ? · move_l ? · adv_to_mark_l ? (is_marked ?).
-
-definition R_init_match ≝ λt1,t2.
- ∀ls,l,rs,c,d. no_grids (〈c,false〉::l) → no_marks l →
- t1 = midtape STape ls 〈c,false〉 (l@〈grid,false〉::〈bar,false〉::〈d,false〉::rs) →
- t2 = midtape STape ls 〈c,true〉 (l@〈grid,false〉::〈bar,false〉::〈d,true〉::rs).
-
-lemma sem_init_match : Realize ? init_match R_init_match.
-#intape
-cases (sem_seq ????? (sem_mark ?)
- (sem_seq ????? (sem_adv_to_mark_r ? (λc:STape.is_grid (\fst c)))
- (sem_seq ????? (sem_move_r ?)
- (sem_seq ????? (sem_move_r ?)
- (sem_seq ????? (sem_mark ?)
- (sem_seq ????? (sem_move_l ?)
- (sem_adv_to_mark_l ? (is_marked ?))))))) intape)
-#k * #outc * #Hloop #HR
-@(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#ls #l #rs #c #d #Hnogrids #Hnomarks #Hintape
-cases HR -HR
-#ta * whd in ⊢ (%→?); * #Hta #_ lapply (Hta … Hintape) -Hta -Hintape #Hta
-* #tb * whd in ⊢ (%→?); * #_ #Htb cases (Htb … Hta) -Htb -Hta
- [* #Hgridc @False_ind @(absurd … Hgridc) @eqnot_to_noteq
- @(Hnogrids 〈c,false〉) @memb_hd ]
-* * #Hgrdic #Htb #_ lapply (Htb l 〈grid,false〉 (〈bar,false〉::〈d,false〉::rs) (refl …) (refl …) ?)
- [#x #membl @Hnogrids @memb_cons @membl] -Htb #Htb
-* #tc * whd in ⊢ (%→?); * #_ #Htc lapply (Htc … Htb) -Htc -Htb #Htc
-* #td * whd in ⊢ (%→?); * #_ #Htd lapply (Htd … Htc) -Htd -Htc #Htd
-* #te * whd in ⊢ (%→?); * #Hte #_ lapply (Hte … Htd) -Hte -Htd #Hte
-* #tf * whd in ⊢ (%→?); * #_ #Htf lapply (Htf … Hte) -Htf -Hte #Htf
-whd in ⊢ (%→?); * #_ #Htg cases (Htg … Htf) -Htg -Htf
-#_ #Htg cases (Htg (refl …)) -Htg #Htg #_
-lapply (Htg (〈grid,false〉::reverse ? l) 〈c,true〉 ls (refl …) (refl …) ?)
- [#x #membl @Hnomarks @daemon] -Htg #Htg >Htg >reverse_cons >reverse_reverse
- >associative_append %
-qed.
-
-(* init_copy
-
- init_current_on_match; (* no marks in current *)
- move_r;
- adv_to_mark_r;
- adv_mark_r;
-
-*)
-
-definition init_copy ≝
- init_current_on_match · move_r ? ·
- adv_to_mark_r ? (is_marked ?) · adv_mark_r ?.
-
-definition R_init_copy ≝ λt1,t2.
- ∀l1,l2,c,ls,d,rs.
- no_marks l1 → no_grids l1 →
- no_marks l2 → is_grid c = false →
- t1 = midtape STape (l1@〈c,false〉::〈grid,false〉::ls) 〈grid,false〉 (l2@〈comma,true〉::〈d,false〉::rs) →
- t2 = midtape STape (〈comma,false〉::(reverse ? l2)@〈grid,false〉::l1@〈c,true〉::〈grid,false〉::ls) 〈d,true〉 rs.
-
-lemma list_last: ∀A.∀l:list A.
- l = [ ] ∨ ∃a,l1. l = l1@[a].
-#A #l <(reverse_reverse ? l) cases (reverse A l)
- [%1 //
- |#a #l1 %2 @(ex_intro ?? a) @(ex_intro ?? (reverse ? l1)) //
- ]
-qed.
-
-lemma sem_init_copy : Realize ? init_copy R_init_copy.
-#intape
-cases (sem_seq ????? sem_init_current_on_match
- (sem_seq ????? (sem_move_r ?)
- (sem_seq ????? (sem_adv_to_mark_r ? (is_marked ?))
- (sem_adv_mark_r ?))) intape)
-#k * #outc * #Hloop #HR
-@(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#l1 #l2 #c #ls #d #rs #Hl1marks #Hl1grids #Hl2marks #Hc #Hintape
-cases HR -HR
-#ta * whd in ⊢ (%→?); #Hta lapply (Hta … Hl1grids Hc Hintape) -Hta -Hintape #Hta
-* #tb * whd in ⊢ (%→?); * #_ #Htb lapply (Htb … Hta) -Htb -Hta
-generalize in match Hl1marks; -Hl1marks cases (list_last ? l1)
- [#eql1 >eql1 #Hl1marks whd in ⊢ ((???%)→?); whd in ⊢ ((???(????%))→?); #Htb
- * #tc * whd in ⊢ (%→?); * #_ #Htc lapply (Htc … Htb) -Htc -Htb *
- [* whd in ⊢ ((??%?)→?); #Htemp destruct (Htemp)]
- * * #_ #Htc #_ lapply (Htc … (refl …) (refl …) ?)
- [#x #membx @Hl2marks @membx]
- #Htc whd in ⊢ (%→?); * #Houtc #_ cases (Houtc (reverse ? l2@〈grid,false〉::〈c,true〉::〈grid,false〉::ls) comma)
- -Houtc #Houtc lapply (Houtc … Htc) -Houtc -Htc #Houtc #_
- >Houtc %
- |* #c1 * #tl #eql1 >eql1 #Hl1marks >reverse_append >reverse_single
- whd in ⊢ ((???%)→?); whd in ⊢ ((???(????%))→?);
- >associative_append whd in ⊢ ((???(????%))→?); #Htb
- * #tc * whd in ⊢ (%→?); * #_ #Htc lapply (Htc … Htb) -Htc -Htb *
- [* >Hl1marks [#Htemp destruct (Htemp)] @memb_append_l2 @memb_hd]
- * * #_ >append_cons <associative_append #Htc lapply (Htc … (refl …) (refl …) ?)
- [#x #membx cases (memb_append … membx) -membx #membx
- [cases (memb_append … membx) -membx #membx
- [@Hl1marks @memb_append_l1 @daemon
- |>(memb_single … membx) %
- ]
- |@Hl2marks @membx
- ]]
- -Htc #Htc #_ whd in ⊢ (%→?); * #Houtc #_ cases (Houtc (reverse (FinProd FSUnialpha FinBool) ((reverse STape tl@[〈grid,false〉])@l2)
- @c1::〈c,true〉::〈grid,false〉::ls) comma)
- -Houtc #Houtc lapply (Houtc … Htc) -Houtc -Htc #Houtc #_
- >Houtc >reverse_append >reverse_append >reverse_single
- >reverse_reverse >associative_append >associative_append
- >associative_append %
-qed.
-
-(* OLD
-definition init_copy ≝
- seq ? (adv_mark_r ?)
- (seq ? init_current_on_match
- (seq ? (move_r ?)
- (adv_to_mark_r ? (is_marked ?)))).
-
-definition R_init_copy ≝ λt1,t2.
- ∀l1,l2,c,l3,d,rs.
- no_marks l1 → no_grids l1 →
- no_marks l2 → no_grids l2 → is_grid c = false → is_grid d =false →
- t1 = midtape STape (l1@〈grid,false〉::l2@〈c,false〉::〈grid,false〉::l3) 〈comma,true〉 (〈d,false〉::rs) →
- t2 = midtape STape (〈comma,false〉::l1@〈grid,false〉::l2@〈c,true〉::〈grid,false〉::l3) 〈d,true〉 rs.
-
-lemma list_last: ∀A.∀l:list A.
- l = [ ] ∨ ∃a,l1. l = l1@[a].
-#A #l <(reverse_reverse ? l) cases (reverse A l)
- [%1 //
- |#a #l1 %2 @(ex_intro ?? a) @(ex_intro ?? (reverse ? l1)) //
- ]
-qed.
-
-lemma sem_init_copy : Realize ? init_copy R_init_copy.
-#intape
-cases (sem_seq ????? (sem_adv_mark_r ?)
- (sem_seq ????? sem_init_current_on_match
- (sem_seq ????? (sem_move_r ?)
- (sem_adv_to_mark_r ? (is_marked ?)))) intape)
-#k * #outc * #Hloop #HR
-@(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
-#l1 #l2 #c #l3 #d #rs #Hl1marks #Hl1grids #Hl2marks #Hl2grids #Hc #Hd #Hintape
-cases HR -HR
-#ta * whd in ⊢ (%→?); #Hta lapply (Hta … Hintape) -Hta -Hintape #Hta
-* #tb * whd in ⊢ (%→?);
->append_cons #Htb lapply (Htb (〈comma,false〉::l1) l2 c … Hta)
- [@Hd |@Hc |@Hl2grids
- |#x #membx cases (orb_true_l … membx) -membx #membx
- [>(\P membx) // | @Hl1grids @membx]
- ] -Htb #Htb
-* #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htc -Htb
->reverse_append >reverse_cons cases (list_last ? l2)
- [#Hl2 >Hl2 >associative_append whd in ⊢ ((???(??%%%))→?); #Htc
- whd in ⊢ (%→?); #Htd cases (Htd … Htc) -Htd -Htc
- [* whd in ⊢ ((??%?)→?); #Habs destruct (Habs)]
- * #_ #Htf lapply (Htf … (refl …) (refl …) ?)
- [#x >reverse_cons #membx cases (memb_append … membx) -membx #membx
- [@Hl1marks @daemon |>(memb_single … membx) //]
- -Htf
- |#Htf >Htf >reverse_reverse >associative_append %
- ]
- |* #a * #l21 #Heq >Heq >reverse_append >reverse_single
- >associative_append >associative_append >associative_append whd in ⊢ ((???(??%%%))→?); #Htc
- whd in ⊢ (%→?); #Htd cases (Htd … Htc) -Htd -Htc
- [* >Hl2marks [#Habs destruct (Habs) |>Heq @memb_append_l2 @memb_hd]]
- * #_ <associative_append <associative_append #Htf lapply (Htf … (refl …) (refl …) ?)
- [#x >reverse_cons #membx cases (memb_append … membx) -membx #membx
- [cases (memb_append … membx) -membx #membx
- [@Hl2marks >Heq @memb_append_l1 @daemon
- |>(memb_single … membx) //]
- |cases (memb_append … membx) -membx #membx
- [@Hl1marks @daemon |>(memb_single … membx) //]
- ]
- | #Htf >Htf >reverse_append >reverse_reverse
- >reverse_append >reverse_reverse >associative_append
- >reverse_single >associative_append >associative_append
- >associative_append %
- ]
- ]
-qed. *)
-
-definition exec_action ≝
- init_copy · copy · move_r … · move_tape.
-
-definition map_move ≝
- λc,mv.match c with [ null ⇒ None ? | _ ⇒ Some ? 〈c,false,move_of_unialpha mv〉 ].
-
-(* - aggiungere a legal_tape le condizioni
- only_bits ls, rs; bit_or_null c
- - ci vuole un lemma che dimostri
- bit_or_null c1 = true bit_or_null mv = true
- mv ≠ null → c1 ≠ null
- dal fatto che c1 e mv sono contenuti nella table
- *)
-
-definition Pre_exec_action ≝ λt1.
- ∃n,curconfig,ls,rs,c0,c1,s0,s1,table1,newconfig,mv,table2.
- table_TM n (table1@〈comma,false〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2) ∧
- no_marks curconfig ∧ only_bits (curconfig@[〈s0,false〉]) ∧
- only_bits (〈s1,false〉::newconfig) ∧ bit_or_null c1 = true ∧
- |curconfig| = |newconfig| ∧
- legal_tape ls 〈c0,false〉 rs ∧
- t1 = midtape STape (〈c0,false〉::curconfig@〈s0,false〉::〈grid,false〉::ls) 〈grid,false〉
- (table1@〈comma,true〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs).
-
-definition R_exec_action ≝ λt1,t2.
- ∀n,curconfig,ls,rs,c0,c1,s0,s1,table1,newconfig,mv,table2.
- table_TM n (table1@〈comma,false〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2) →
- no_marks curconfig → only_bits (curconfig@[〈s0,false〉]) →
- only_bits (〈s1,false〉::newconfig) → bit_or_null c1 = true →
- |curconfig| = |newconfig| →
- legal_tape ls 〈c0,false〉 rs →
- t1 = midtape STape (〈c0,false〉::curconfig@〈s0,false〉::〈grid,false〉::ls) 〈grid,false〉
- (table1@〈comma,true〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs) →
- ∀t1'.t1' = lift_tape ls 〈c0,false〉 rs →
- ∃ls1,rs1,c2.
- t2 = midtape STape ls1 〈grid,false〉
- (〈s1,false〉::newconfig@〈c2,false〉::〈grid,false〉::
- table1@〈comma,false〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs1) ∧
- lift_tape ls1 〈c2,false〉 rs1 =
- tape_move STape t1' (map_move c1 mv) ∧ legal_tape ls1 〈c2,false〉 rs1.
-
-(* move the following 2 lemmata to mono.ma *)
-lemma tape_move_left_eq :
- ∀A.∀t:tape A.∀c.
- tape_move ? t (Some ? 〈c,L〉) =
- tape_move_left ? (left ? t) c (right ? t).
-//
-qed.
-
-lemma tape_move_right_eq :
- ∀A.∀t:tape A.∀c.
- tape_move ? t (Some ? 〈c,R〉) =
- tape_move_right ? (left ? t) c (right ? t).
-//
-qed.
-
-lemma lift_tape_not_null :
- ∀ls,c,bc,rs.c ≠ null → lift_tape ls 〈c,bc〉 rs = midtape ? ls 〈c,bc〉 rs.
-#ls #c #bc #rs cases c //
-#Hfalse @False_ind /2/
-qed.
-
-lemma merge_char_not_null :
- ∀c1,c2.c1 ≠ null → merge_char c1 c2 ≠ null.
-#c1 #c2 @not_to_not cases c2
-[ #c1' normalize #Hfalse destruct (Hfalse)
-| normalize //
-| *: normalize #Hfalse destruct (Hfalse)
-]
-qed.
-
-lemma merge_char_null : ∀c.merge_char null c = c.
-* //
-qed.
-
-lemma merge_char_cases : ∀c1,c2.merge_char c1 c2 = c1 ∨ merge_char c1 c2 = c2.
-#c1 *
-[ #c1' %2 %
-| % %
-| *: %2 % ]
-qed.
-
-(* lemma merge_char_c_bit :
- ∀c1,c2.is_bit c2 = true → merge_char c1 c2 = c2.
-#c1 *
-[ #c2' #_ %
-|*: normalize #Hfalse destruct (Hfalse) ]
-qed.
-
-lemma merge_char_c_bit :
- ∀c1,c2.is_null c2 = true → merge_char c1 c2 = c1.
-#c1 *
-[ #c2' #_ %
-|*: normalize #Hfalse destruct (Hfalse) ]
-qed.
-
-*)
-
-(*lemma GRealize_to_Realize :
- ∀sig,M,R.GRealize sig M (λx.True) R → Realize sig M R.
-#sig #M #R #HR #t @HR //
-qed.*)
-
-lemma sem_exec_action : GRealize ? exec_action Pre_exec_action R_exec_action.
-@(sem_seq_app_guarded … (Realize_to_GRealize … sem_init_copy) ???)
-[3:@(sem_seq_app_guarded … sem_copy)
- [3:@(sem_seq_app_guarded … (Realize_to_GRealize … (sem_move_r STape)))
- [3:@(Realize_to_GRealize … (sem_move_tape …))
- |@(λt.True)
- |4://
- |5:@sub_reflexive]
- |@(λt.True)
- |4://
- |5:@sub_reflexive]
-|4:#t1 #t2
- * #n * #curconfig * #ls * #rs * #c0 * #c1 * #s0 * #s1 * #table1 * #newconfig
- * #mv * #table2 * * * * * * *
- #Htable #Hcurconfig1 #Hcurconfig2 #Hnewconfig #Hc1 #Hlen #Htape #Ht1
- whd in ⊢ (%→?); >Ht1 #HR
- lapply (HR (〈c0,false〉::curconfig) table1 s0 ls s1
- (newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs) ???? (refl ??))
- [(*Hcurconfig2*) @daemon
- |(*by Htable *) @daemon
- |(*Hcurconfig2*) @daemon
- |(*Hcurconfig1*) @daemon ]
- -HR #Ht2 whd
- %{(〈grid,false〉::ls)} %{s0} %{s1} %{c0} %{c1} %{(〈mv,false〉::table2@〈grid,false〉::rs)}
- %{newconfig} %{(〈comma,false〉::reverse ? table1)} %{curconfig} >Ht2
- % [ % [ % [ % [ % [ % [ % [ %
- [ %
- |(*by Htabel*) @daemon ]
- |(*by Htable*) @daemon ]
- |// ]
- |>Hlen % ]
- |@Hcurconfig2 ]
- |@Hnewconfig ]
- |cases Htape * * // ]
- | // ]
- |1,2:skip]
-#ta #outtape #HR #n #curconfig #ls #rs #c0 #c1 #s0 #s1 #table1 #newconfig #mv
-#table2 #Htable #Hcurconfig1 #Hcurconfig2 #Hnewconfig #Hbitnullc1 #Hlen #Htape
-#Hta #ta' #Hta'
-cases HR -HR #tb * whd in ⊢ (%→?); #Htb
-lapply (Htb (〈c0,false〉::curconfig) table1 s0 ls s1
- (newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs)
- ???? Hta) -Htb
-[ (*Hcurconfig2*) @daemon
-| (*Htable*) @daemon
-| (* by Htape bit_or_null c0 = true, moreover Hcurconfig2 *) @daemon
-| (*Hcurconfig1*) @daemon ]
-#Htb * #tc * whd in ⊢ (%→?); #Htc
-lapply (Htc (〈grid,false〉::ls) s0 s1 c0 c1 (〈mv,false〉::table2@〈grid,false〉::rs) newconfig (〈comma,false〉::reverse ? table1) curconfig Htb ????????) -Htc
-[9:|*:(* bit_or_null c0,c1; |curconfig| = |newconfig|*) @daemon ]
-#Htc * #td * whd in ⊢ (%→?); * #_ #Htd lapply (Htd … Htc) -Htd whd in ⊢(???(??%%%)→?);#Htd
-whd in ⊢ (%→?); #Houtc whd in Htd:(???%); whd in Htd:(???(??%%%));
-lapply (Houtc rs n
- (〈comma,false〉::〈c1,false〉::reverse ? newconfig@〈s1,false〉::〈comma,false〉::reverse ? table1)
- mv table2 (merge_char c0 c1) (reverse ? newconfig@[〈s1,false〉]) ls ????????)
-[3: cases Htape -Htape * * #Hnomarks #Hbits #Hc0 #Hlsrs % [ % [ %
- [ #x #Hx cases (orb_true_l … Hx) #Hx'
- [ >(\P Hx') %
- | @Hnomarks @memb_cons // ]
- | @Hbits ]
- | cases (merge_char_cases c0 c1) #Hmerge >Hmerge // ]
- | cases (true_or_false (c0 == null)) #Hc0'
- [ cases Hlsrs -Hlsrs
- [ *
- [ >(\P Hc0') * #Hfalse @False_ind /2/
- | #Hlsnil % %2 // ]
- | #Hrsnil %2 // ]
- | % % @merge_char_not_null @(\Pf Hc0') ] ]
-|4:>Htd @(eq_f3 … (midtape ?))
- [ @eq_f @eq_f >associative_append >associative_append %
- | %
- | % ]
-| %
-|| >reverse_cons >reverse_cons >reverse_append >reverse_reverse
- >reverse_cons >reverse_cons >reverse_reverse
- >associative_append >associative_append >associative_append
- >associative_append >associative_append
- @Htable
-| (* well formedness of table *) @daemon
-| (* Hnewconfig *) @daemon
-| (* bit_or_null mv = true (well formedness of table) *) @daemon
-| -Houtc * #ls1 * #rs1 * #newc * #Hnewtapelegal * #Houtc *
- [ *
- [ * #Hmv #Htapemove
- @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc)
- %
- [ %
- [ >Houtc -Houtc >reverse_append
- >reverse_reverse >reverse_single @eq_f
- >reverse_cons >reverse_cons >reverse_append >reverse_cons
- >reverse_cons >reverse_reverse >reverse_reverse
- >associative_append >associative_append
- >associative_append >associative_append
- >associative_append >associative_append %
- | >Hmv >Hta' >Htapemove
- (* mv = bit false -→ c1 = bit ? *)
- cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1
- >Hc1 >tape_move_left_eq >(legal_tape_left … Htape)
- >(legal_tape_right … Htape) %
- ]
- | //
- ]
- | * #Hmv #Htapemove
- @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) %
- [ %
- [ >Houtc -Houtc >reverse_append
- >reverse_reverse >reverse_single @eq_f
- >reverse_cons >reverse_cons >reverse_append >reverse_cons
- >reverse_cons >reverse_reverse >reverse_reverse
- >associative_append >associative_append
- >associative_append >associative_append
- >associative_append >associative_append %
- |>Hmv >Hta' >Htapemove
- cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1
- >Hc1 >tape_move_right_eq >(legal_tape_left … Htape)
- >(legal_tape_right … Htape) %
- ]
- | //
- ]
- ]
- | * * * #Hmv #Hlseq #Hrseq #Hnewc
- @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) %
- [ %
- [ >Houtc -Houtc >reverse_append
- >reverse_reverse >reverse_single @eq_f
- >reverse_cons >reverse_cons >reverse_append >reverse_cons
- >reverse_cons >reverse_reverse >reverse_reverse
- >associative_append >associative_append
- >associative_append >associative_append
- >associative_append >associative_append %
- |>Hmv >Hta' cases c1 in Hnewc;
- [ #c1' whd in ⊢ (??%?→?);#Hnewc <Hnewc
- >Hlseq >Hrseq whd in ⊢ (??%%);
- >(legal_tape_left … Htape) >(legal_tape_right … Htape) %
- | whd in ⊢ (??%?→?); #Hnewc >Hnewc >Hlseq >Hrseq %
- |*: whd in ⊢ (??%?→?);#Hnewc <Hnewc
- >Hlseq >Hrseq whd in ⊢ (??%%);
- >(legal_tape_left … Htape) >(legal_tape_right … Htape) %
- ]
- ]
- | //
- ]
- ]
-]
-qed.
-
-(*
-if is_false(current) (* current state is not final *)
- then init_match;
- match_tuple;
- if is_marked(current) = false (* match ok *)
- then
- exec_action
- move_r;
- else sink;
- else nop;
-*)
-
-definition uni_step ≝
- ifTM ? (test_char STape (λc.\fst c == bit false))
- (single_finalTM ?
- (init_match · match_tuple ·
- (ifTM ? (test_char ? (λc.¬is_marked ? c))
- (exec_action · move_r …)
- (nop ?) tc_true)))
- (nop ?) tc_true.
-
-definition R_uni_step_true ≝ λt1,t2.
- ∀n,table,s0,s1,c0,c1,ls,rs,curconfig,newconfig,mv.
- 0 < |table| → table_TM (S n) table →
- match_in_table (S n) (〈s0,false〉::curconfig) 〈c0,false〉
- (〈s1,false〉::newconfig) 〈c1,false〉 〈mv,false〉 table →
- legal_tape ls 〈c0,false〉 rs →
- t1 = midtape STape (〈grid,false〉::ls) 〈s0,false〉
- (curconfig@〈c0,false〉::〈grid,false〉::table@〈grid,false〉::rs) →
- ∀t1'.t1' = lift_tape ls 〈c0,false〉 rs →
- s0 = bit false ∧
- ∃ls1,rs1,c2.
- (t2 = midtape STape (〈grid,false〉::ls1) 〈s1,false〉
- (newconfig@〈c2,false〉::〈grid,false〉::table@〈grid,false〉::rs1) ∧
- lift_tape ls1 〈c2,false〉 rs1 =
- tape_move STape t1' (map_move c1 mv) ∧ legal_tape ls1 〈c2,false〉 rs1).
-
-definition R_uni_step_false ≝ λt1,t2.
- ∀b. current STape t1 = Some ? 〈bit b,false〉 → b = true ∧ t2 = t1.
-
-(*axiom sem_match_tuple : Realize ? match_tuple R_match_tuple.*)
-
-definition us_acc : states ? uni_step ≝ (inr … (inl … (inr … start_nop))).
-
-definition Pre_uni_step ≝ λt1.
- ∃n,table,s0,s1,c0,c1,ls,rs,curconfig,newconfig,mv.
- 0 < |table| ∧ table_TM (S n) table ∧
- match_in_table (S n) (〈s0,false〉::curconfig) 〈c0,false〉
- (〈s1,false〉::newconfig) 〈c1,false〉 〈mv,false〉 table ∧
- legal_tape ls 〈c0,false〉 rs ∧
- t1 = midtape STape (〈grid,false〉::ls) 〈s0,false〉
- (curconfig@〈c0,false〉::〈grid,false〉::table@〈grid,false〉::rs).
-
-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) ???)
- [ 5: @sub_reflexive
- | 3: @(sem_seq_app_guarded … sem_match_tuple
- (Realize_to_GRealize … (sem_if ????????? (sem_test_char … (λc.¬is_marked FSUnialpha c))
- (sem_seq … sem_exec_action (sem_move_r …))
- (sem_nop …))))
- [@(λx.True)
- |//
- |@sub_reflexive]
- ||| #t1 #t2 * #n * #table * #s0 * #s1 * #c0 * #c1 * #ls * #rs * #curconfig
- * #newconfig * #mv * * * *
- #Hlen1 #Htable #Hmatch #Hlegal #Ht1
- whd in ⊢ (%→?);
- cut (∃tup,table0.table = tup@table0 ∧ tuple_TM (S n) tup)
- [@daemon]
- * #tup * #table0 * #Htableeq * #qin * #cin * #qout * #cout * #mv0
- * * * * * * * * * *
- #Hqinnomarks #_ #Hqinbits #_ #_ #_ #_ #_ #Hqinlen #_ #Htupeq
- cut (∃d,qin0.qin = 〈d,false〉::qin0)
- [ lapply Hqinlen lapply Hqinnomarks -Hqinlen -Hqinnomarks cases qin
- [ #_ normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
- | * #d #bd #qin0 #Hqinnomarks #_ %{d} %{qin0}
- >(?:bd=false) [%]
- @(Hqinnomarks 〈d,bd〉) @memb_hd ] ]
- * #d * #qin0 #Hqineq
- #Ht2
- lapply (Ht2 (〈grid,false〉::ls) (curconfig@[〈c0,false〉])
- (qin0@〈cin,false〉::〈comma,false〉::qout@〈cout,false〉::〈comma,false〉::〈mv0,false〉::table0@〈grid,false〉::rs) s0 d ???)
- [ >Ht1 @eq_f >associative_append @eq_f @eq_f @eq_f
- >Htableeq >Htupeq >associative_append whd in ⊢ (??%?);
- @eq_f >Hqineq >associative_append @eq_f whd in ⊢ (??%?);
- @eq_f whd in ⊢ (??%?); @eq_f
- >associative_append %
- | @daemon
- | @daemon
- ]
- #Ht2 % [| % [| % [| % [ @Ht2 ]
- %2
- (* ls0 = ls
- c = s0
- l1 = curconfig@[〈c0,false〉]
- l2 = [〈bar,false〉]
- c10 = d
- l3 = qin0@[〈cin,false〉]
- l4 = qout@〈cout,false〉::〈comma,false〉::〈mv0,false〉::table0
- rs00 = rs
- n0 = S n ?
- *)
- %{ls} %{s0} %{(curconfig@[〈c0,false〉])}
- %{([〈bar,false〉])} %{d} %{(qin0@[〈cin,false〉])}
- %{(qout@〈cout,false〉::〈comma,false〉::〈mv0,false〉::table0)}
- %{rs} %{n} @daemon (* TODO *)
- ]
- ]
- ]
- ]
- | #intape #outtape
- #ta whd in ⊢ (%→?); #Hta #HR
- #n #fulltable #s0 #s1 #c0 #c1 #ls #rs #curconfig #newconfig #mv
- #Htable_len cut (∃t0,table. fulltable =〈bar,false〉::〈t0,false〉::table) [(* 0 < |table| *) @daemon]
- * #t0 * #table #Hfulltable >Hfulltable -fulltable
- #Htable #Hmatch #Htape #Hintape #t1' #Ht1'
- >Hintape in Hta; * * * #c #bc *
- whd in ⊢ (??%?→?); #HSome destruct (HSome) #Hc #Hta % [@(\P Hc)]
- cases HR -HR
- #tb * whd in ⊢ (%→?); #Htb
- lapply (Htb (〈grid,false〉::ls) (curconfig@[〈c0,false〉]) (table@〈grid,false〉::rs) c t0 ???)
- [ >Hta >associative_append %
- | @daemon
- | @daemon
- | -Hta -Htb #Htb *
- #tc * whd in ⊢ (%→?); #Htc cases (Htc … Htable … Htb) -Htb -Htc
- [| * #Hcurrent #Hfalse @False_ind
- (* absurd by Hmatch *) @daemon
- | >(\P Hc) %
- | (* Htable (con lemma) *) @daemon
- | (* Hmatch *) @daemon
- | (* Htable *) @daemon
- | (* Htable, Hmatch → |config| = n
- necessaria modifica in R_match_tuple, le dimensioni non corrispondono
- *) @daemon
- ]
- * #table1 * #newc * #mv1 * #table2 * #Htableeq #Htc *
- [ * #td * whd in ⊢ (%→?); >Htc -Htc * * #c2 * whd in ⊢ (??%?→?); #Hc2 destruct (Hc2)
- #_ #Htd
- cut (newc = 〈s1,false〉::newconfig@[〈c1,false〉]) [@daemon] #Hnewc
- >Hnewc cut (mv1 = 〈mv,false〉)
- [@daemon] #Hmv1
- * #te * whd in ⊢ (%→?); #Hte
- cut (td = midtape STape (〈c0,false〉::reverse STape curconfig@〈c,false〉::〈grid,false〉::ls)
- 〈grid,false〉
- ((table1@〈bar,false〉::〈c,false〉::curconfig@[〈c0,false〉])@〈comma,true〉::〈s1,false〉::
- newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs))
- [ >Htd @eq_f3 //
- [ >reverse_append >reverse_single %
- | >associative_append >associative_append normalize
- >associative_append >Hmv1 >Hnewc @eq_f @eq_f @eq_f @eq_f @eq_f @eq_f
- whd in ⊢ (??%?); >associative_append %
- ]
- ]
- -Htd #Htd lapply (Hte … (S n) … Htd … Ht1') -Htd -Hte
- [ //
- | (*|curconfig| = |newconfig|*) @daemon
- | (* Htable → bit_or_null c1 = true *) @daemon
- | (* only_bits (〈s1,false〉::newconfig) *) @daemon
- | (* only_bits (curconfig@[〈s0,false〉]) *) @daemon
- | (* no_marks (reverse ? curconfig) *) @daemon
- | >Hmv1 in Htableeq; >Hnewc
- >associative_append >associative_append normalize
- >associative_append >associative_append
- #Htableeq <Htableeq // ]
- * #ls1 * #rs1 * #c2 * * #Hte #Hliftte #Hlegalte
- whd in ⊢ (%→?); * #_ #Houttape lapply (Houttape … Hte) -Houttape #Houttape
- whd in Houttape:(???%); whd in Houttape:(???(??%%%));
- @ex_intro [| @(ex_intro ?? rs1) @ex_intro [| % [ %
- [ >Houttape @eq_f @eq_f @eq_f @eq_f
- change with ((〈bar,false〉::〈t0,false〉::table)@?) in ⊢ (???%);
- >Htableeq >associative_append >associative_append
- >associative_append normalize >associative_append
- >associative_append normalize >Hnewc <Hmv1
- >associative_append normalize >associative_append
- >Hmv1 %
- | @Hliftte
- ]
- | //
- ]
- ]
- ]
- | * #td * whd in ⊢ (%→%→?); >Htc * #Htd
- lapply (Htd ? (refl ??)) normalize in ⊢ (%→?);
- #Hfalse destruct (Hfalse)
- ]
- ]
-| #t1 #t2 #t3 whd in ⊢ (%→%→?); #Ht1 #Ht2
- #b #Hb >Hb in Ht1; * #Hc #Ht1 lapply (Hc ? (refl ??)) -Hc #Hb' %
- // cases b in Hb'; normalize #H1 //
-]
-*)
-qed.
-
-(*
-@(acc_sem_if_app STape … (sem_test_char ? (λc:STape.\fst c == bit false))
- (sem_seq … sem_init_match
- (sem_seq … sem_match_tuple
- (sem_if … (* ????????? (sem_test_char … (λc.¬is_marked FSUnialpha c)) *)
- (sem_seq … sem_exec_action (sem_move_r …))
- (sem_nop …))))
- (sem_nop …)
- …)
-[@sem_test_char||]
-[ #intape #outtape
- #ta whd in ⊢ (%→?); #Hta #HR
- #n #fulltable #s0 #s1 #c0 #c1 #ls #rs #curconfig #newconfig #mv
- #Htable_len cut (∃t0,table. fulltable =〈bar,false〉::〈t0,false〉::table) [(* 0 < |table| *) @daemon]
- * #t0 * #table #Hfulltable >Hfulltable -fulltable
- #Htable #Hmatch #Htape #Hintape #t1' #Ht1'
- >Hintape in Hta; * * * #c #bc *
- whd in ⊢ (??%?→?); #HSome destruct (HSome) #Hc #Hta % [@(\P Hc)]
- cases HR -HR
- #tb * whd in ⊢ (%→?); #Htb
- lapply (Htb (〈grid,false〉::ls) (curconfig@[〈c0,false〉]) (table@〈grid,false〉::rs) c t0 ???)
- [ >Hta >associative_append %
- | @daemon
- | @daemon
- | -Hta -Htb #Htb *
- #tc * whd in ⊢ (%→?); #Htc cases (Htc … Htable … Htb) -Htb -Htc
- [| * #Hcurrent #Hfalse @False_ind
- (* absurd by Hmatch *) @daemon
- | >(\P Hc) %
- | (* Htable (con lemma) *) @daemon
- | (* Hmatch *) @daemon
- | (* Htable *) @daemon
- | (* Htable, Hmatch → |config| = n
- necessaria modifica in R_match_tuple, le dimensioni non corrispondono
- *) @daemon
- ]
- * #table1 * #newc * #mv1 * #table2 * #Htableeq #Htc *
- [ * #td * whd in ⊢ (%→?); >Htc -Htc * * #c2 * whd in ⊢ (??%?→?); #Hc2 destruct (Hc2)
- #_ #Htd
- cut (newc = 〈s1,false〉::newconfig@[〈c1,false〉]) [@daemon] #Hnewc
- >Hnewc cut (mv1 = 〈mv,false〉)
- [@daemon] #Hmv1
- * #te * whd in ⊢ (%→?); #Hte
- cut (td = midtape STape (〈c0,false〉::reverse STape curconfig@〈c,false〉::〈grid,false〉::ls)
- 〈grid,false〉
- ((table1@〈bar,false〉::〈c,false〉::curconfig@[〈c0,false〉])@〈comma,true〉::〈s1,false〉::
- newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs))
- [ >Htd @eq_f3 //
- [ >reverse_append >reverse_single %
- | >associative_append >associative_append normalize
- >associative_append >Hmv1 >Hnewc @eq_f @eq_f @eq_f @eq_f @eq_f @eq_f
- whd in ⊢ (??%?); >associative_append %
- ]
- ]
- -Htd #Htd lapply (Hte … (S n) … Htd … Ht1') -Htd -Hte
- [ //
- | (*|curconfig| = |newconfig|*) @daemon
- | (* Htable → bit_or_null c1 = true *) @daemon
- | (* only_bits (〈s1,false〉::newconfig) *) @daemon
- | (* only_bits (curconfig@[〈s0,false〉]) *) @daemon
- | (* no_marks (reverse ? curconfig) *) @daemon
- | >Hmv1 in Htableeq; >Hnewc
- >associative_append >associative_append normalize
- >associative_append >associative_append
- #Htableeq <Htableeq // ]
- * #ls1 * #rs1 * #c2 * * #Hte #Hliftte #Hlegalte
- whd in ⊢ (%→?); * #_ #Houttape lapply (Houttape … Hte) -Houttape #Houttape
- whd in Houttape:(???%); whd in Houttape:(???(??%%%));
- @ex_intro [| @(ex_intro ?? rs1) @ex_intro [| % [ %
- [ >Houttape @eq_f @eq_f @eq_f @eq_f
- change with ((〈bar,false〉::〈t0,false〉::table)@?) in ⊢ (???%);
- >Htableeq >associative_append >associative_append
- >associative_append normalize >associative_append
- >associative_append normalize >Hnewc <Hmv1
- >associative_append normalize >associative_append
- >Hmv1 %
- | @Hliftte
- ]
- | //
- ]
- ]
- ]
- | * #td * whd in ⊢ (%→%→?); >Htc * #Htd
- lapply (Htd ? (refl ??)) normalize in ⊢ (%→?);
- #Hfalse destruct (Hfalse)
- ]
- ]
-| #t1 #t2 #t3 whd in ⊢ (%→%→?); #Ht1 #Ht2
- #b #Hb >Hb in Ht1; * #Hc #Ht1 lapply (Hc ? (refl ??)) -Hc #Hb' %
- // cases b in Hb'; normalize #H1 //
-]
-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/universal/uni_step.ma".
-
-(* definition zero : ∀n.initN n ≝ λn.mk_Sig ?? 0 (le_O_n n). *)
-
-definition low_config: ∀M:normalTM.nconfig (no_states M) → tape STape ≝
-λM:normalTM.λc.
- let n ≝ no_states M in
- let h ≝ nhalt M in
- let t ≝ntrans M in
- let q ≝ cstate … c in
- let q_low ≝ m_bits_of_state n h q in
- let current_low ≝ match current … (ctape … c) with [ None ⇒ null | Some b ⇒ bit b] in
- let low_left ≝ map … (λb.〈bit b,false〉) (left … (ctape …c)) in
- let low_right ≝ map … (λb.〈bit b,false〉) (right … (ctape …c)) in
- let table ≝ flatten ? (tuples_list n h (graph_enum ?? t)) in
- let right ≝ q_low@〈current_low,false〉::〈grid,false〉::table@〈grid,false〉::low_right in
- mk_tape STape (〈grid,false〉::low_left) (option_hd … right) (tail … right).
-
-lemma low_config_eq: ∀t,M,c. t = low_config M c →
- ∃low_left,low_right,table,q_low_hd,q_low_tl,c_low.
- low_left = map … (λb.〈bit b,false〉) (left … (ctape …c)) ∧
- low_right = map … (λb.〈bit b,false〉) (right … (ctape …c)) ∧
- table = flatten ? (tuples_list (no_states M) (nhalt M) (graph_enum ?? (ntrans M))) ∧
- 〈q_low_hd,false〉::q_low_tl = m_bits_of_state (no_states M) (nhalt M) (cstate … c) ∧
- c_low = match current … (ctape … c) with [ None ⇒ null| Some b ⇒ bit b] ∧
- t = midtape STape (〈grid,false〉::low_left) 〈q_low_hd,false〉 (q_low_tl@〈c_low,false〉::〈grid,false〉::table@〈grid,false〉::low_right).
-#t #M #c #eqt
- @(ex_intro … (map … (λb.〈bit b,false〉) (left … (ctape …c))))
- @(ex_intro … (map … (λb.〈bit b,false〉) (right … (ctape …c))))
- @(ex_intro … (flatten ? (tuples_list (no_states M) (nhalt M) (graph_enum ?? (ntrans M)))))
- @(ex_intro … (\fst (hd ? (m_bits_of_state (no_states M) (nhalt M) (cstate … c)) 〈bit true,false〉)))
- @(ex_intro … (tail ? (m_bits_of_state (no_states M) (nhalt M) (cstate … c))))
- @(ex_intro … (match current … (ctape … c) with [ None ⇒ null| Some b ⇒ bit b]))
-% [% [% [% [% // | // ] | // ] | // ] | >eqt //]
-qed.
-
-let rec to_bool_list (l: list (unialpha×bool)) ≝
- match l with
- [ nil ⇒ nil ?
- | cons a tl ⇒
- match \fst a with
- [bit b ⇒ b::to_bool_list tl
- |_ ⇒ nil ?
- ]
- ].
-
-definition high_c ≝ λc:unialpha×bool.
- match \fst c with
- [ null ⇒ None ?
- | bit b ⇒ Some ? b
- | _ ⇒ None ?].
-
-definition high_tape ≝ λls,c,rs.
- mk_tape FinBool (to_bool_list ls) (high_c c) (to_bool_list rs).
-
-lemma high_tape_eq : ∀ls,c,rs. high_tape ls c rs =
- mk_tape FinBool (to_bool_list ls) (high_c c) (to_bool_list rs).
-// qed.
-
-definition high_tape_from_tape ≝ λt:tape STape.
- match t with
- [niltape ⇒ niltape ?
- |leftof a l ⇒ match \fst a with
- [bit b ⇒ leftof ? b (to_bool_list l)
- |_ ⇒ niltape ?
- ]
- |rightof a r ⇒ match \fst a with
- [bit b ⇒ rightof ? b (to_bool_list r)
- |_ ⇒ niltape ?
- ]
- |midtape l c r ⇒ high_tape l c r
- ].
-
-lemma high_tape_of_lift : ∀ls,c,rs. legal_tape ls c rs →
- high_tape ls c rs =
- high_tape_from_tape (lift_tape ls c rs).
-#ls * #c #b #rs * #H cases c //
->high_tape_eq
-* [ * [#H @False_ind /2/
- | #Heq >Heq cases rs // * #a #b1 #tl
- whd in match (lift_tape ???); cases a //
- ]
- |#Heq >Heq cases ls // * #a #b1 #tl
- whd in match (lift_tape ???); cases a //
- ]
-qed.
-
-lemma bool_embedding: ∀l.
- to_bool_list (map ?? (λb.〈bit b,false〉) l) = l.
-#l elim l // #a #tl #Hind normalize @eq_f @Hind
-qed.
-
-lemma current_embedding: ∀c.
- high_c (〈match c with [None ⇒ null | Some b ⇒ bit b],false〉) = c.
- * normalize // qed.
-
-lemma tape_embedding: ∀ls,c,rs.
- high_tape
- (map ?? (λb.〈bit b,false〉) ls)
- (〈match c with [None ⇒ null | Some b ⇒ bit b],false〉)
- (map ?? (λb.〈bit b,false〉) rs) = mk_tape ? ls c rs.
-#ls #c #rs >high_tape_eq >bool_embedding >bool_embedding
->current_embedding %
-qed.
-
-definition high_move ≝ λc,mv.
- match c with
- [ bit b ⇒ Some ? 〈b,move_of_unialpha mv〉
- | _ ⇒ None ?
- ].
-
-definition map_move ≝
- λc,mv.match c with [ null ⇒ None ? | _ ⇒ Some ? 〈c,false,move_of_unialpha mv〉 ].
-
-definition low_step_R_true ≝ λt1,t2.
- ∀M:normalTM.
- ∀c: nconfig (no_states M).
- t1 = low_config M c →
- halt ? M (cstate … c) = false ∧
- t2 = low_config M (step ? M c).
-
-definition low_tape_aux : ∀M:normalTM.tape FinBool → tape STape ≝
-λM:normalTM.λt.
- let current_low ≝ match current … t with
- [ None ⇒ None ? | Some b ⇒ Some ? 〈bit b,false〉] in
- let low_left ≝ map … (λb.〈bit b,false〉) (left … t) in
- let low_right ≝ map … (λb.〈bit b,false〉) (right … t) in
- mk_tape STape low_left current_low low_right.
-
-lemma left_of_low_tape: ∀M,t.
- left ? (low_tape_aux M t) = map … (λb.〈bit b,false〉) (left … t).
-#M * //
-qed.
-
-lemma right_of_low_tape: ∀M,t.
- right ? (low_tape_aux M t) = map … (λb.〈bit b,false〉) (right … t).
-#M * //
-qed.
-
-definition low_move ≝ λaction:option (bool × move).
- match action with
- [None ⇒ None ?
- |Some act ⇒ Some ? (〈〈bit (\fst act),false〉,\snd act〉)].
-
-(* simulation lemma *)
-lemma low_tape_move : ∀M,action,t.
- tape_move STape (low_tape_aux M t) (low_move action) =
- low_tape_aux M (tape_move FinBool t action).
-#M * // (* None *)
-* #b #mv #t cases mv cases t //
- [#ls #c #rs cases ls //|#ls #c #rs cases rs //]
-qed.
-
-lemma left_of_lift: ∀ls,c,rs. left ? (lift_tape ls c rs) = ls.
-#ls * #c #b #rs cases c // cases ls // cases rs //
-qed.
-
-lemma right_of_lift: ∀ls,c,rs. legal_tape ls c rs →
- right ? (lift_tape ls c rs) = rs.
-#ls * #c #b #rs * #_ cases c // cases ls cases rs // #a #tll #b #tlr
-#H @False_ind cases H [* [#H1 /2/ |#H1 destruct] |#H1 destruct]
-qed.
-
-
-lemma current_of_lift: ∀ls,c,b,rs. legal_tape ls 〈c,b〉 rs →
- current STape (lift_tape ls 〈c,b〉 rs) =
- match c with [null ⇒ None ? | _ ⇒ Some ? 〈c,b〉].
-#ls #c #b #rs cases c // whd in ⊢ (%→?); * #_
-* [* [#Hnull @False_ind /2/ | #Hls >Hls whd in ⊢ (??%%); cases rs //]
- |#Hrs >Hrs whd in ⊢ (??%%); cases ls //]
-qed.
-
-lemma current_of_lift_None: ∀ls,c,b,rs. legal_tape ls 〈c,b〉 rs →
- current STape (lift_tape ls 〈c,b〉 rs) = None ? →
- c = null.
-#ls #c #b #rs #Hlegal >(current_of_lift … Hlegal) cases c normalize
- [#b #H destruct |// |3,4,5:#H destruct ]
-qed.
-
-lemma current_of_lift_Some: ∀ls,c,c1,rs. legal_tape ls c rs →
- current STape (lift_tape ls c rs) = Some ? c1 →
- c = c1.
-#ls * #c #cb #b #rs #Hlegal >(current_of_lift … Hlegal) cases c normalize
- [#b1 #H destruct // |#H destruct |3,4,5:#H destruct //]
-qed.
-
-lemma current_of_low_None: ∀M,t. current FinBool t = None ? →
- current STape (low_tape_aux M t) = None ?.
-#M #t cases t // #l #b #r whd in ⊢ ((??%?)→?); #H destruct
-qed.
-
-lemma current_of_low_Some: ∀M,t,b. current FinBool t = Some ? b →
- current STape (low_tape_aux M t) = Some ? 〈bit b,false〉.
-#M #t cases t
- [#b whd in ⊢ ((??%?)→?); #H destruct
- |#b #l #b1 whd in ⊢ ((??%?)→?); #H destruct
- |#b #l #b1 whd in ⊢ ((??%?)→?); #H destruct
- |#c #c1 #l #r whd in ⊢ ((??%?)→?); #H destruct %
- ]
-qed.
-
-lemma current_of_low:∀M,tape,ls,c,rs. legal_tape ls c rs →
- lift_tape ls c rs = low_tape_aux M tape →
- c = 〈match current … tape with
- [ None ⇒ null | Some b ⇒ bit b], false〉.
-#M #tape #ls * #c #cb #rs #Hlegal #Hlift
-cut (current ? (lift_tape ls 〈c,cb〉 rs) = current ? (low_tape_aux M tape))
- [@eq_f @Hlift] -Hlift #Hlift
-cut (current … tape = None ? ∨ ∃b.current … tape = Some ? b)
- [cases (current … tape) [%1 // | #b1 %2 /2/ ]] *
- [#Hcurrent >Hcurrent normalize
- >(current_of_low_None …Hcurrent) in Hlift; #Hlift
- >(current_of_lift_None … Hlegal Hlift)
- @eq_f cases Hlegal * * #Hmarks #_ #_ #_ @(Hmarks 〈c,cb〉) @memb_hd
- |* #b #Hcurrent >Hcurrent normalize
- >(current_of_low_Some …Hcurrent) in Hlift; #Hlift
- @(current_of_lift_Some … Hlegal Hlift)
- ]
-qed.
-
-(*
-lemma current_of_low:∀M,tape,ls,c,rs. legal_tape ls c rs →
- lift_tape ls c rs = low_tape_aux M tape →
- c = 〈match current … tape with
- [ None ⇒ null | Some b ⇒ bit b], false〉.
-#M #tape #ls * #c #cb #rs * * #_ #H cases (orb_true_l … H)
- [cases c [2,3,4,5: whd in ⊢ ((??%?)→?); #Hfalse destruct]
- #b #_ #_ cases tape
- [whd in ⊢ ((??%%)→?); #H destruct
- |#a #l whd in ⊢ ((??%%)→?); #H destruct
- |#a #l whd in ⊢ ((??%%)→?); #H destruct
- |#a #l #r whd in ⊢ ((??%%)→?); #H destruct //
- ]
- |cases c
- [#b whd in ⊢ ((??%?)→?); #Hfalse destruct
- |3,4,5:whd in ⊢ ((??%?)→?); #Hfalse destruct]
- #_ * [* [#Habs @False_ind /2/
- |#Hls >Hls whd in ⊢ ((??%%)→?); *)
-
-
-(* sufficent conditions to have a low_level_config *)
-lemma is_low_config: ∀ls,c,rs,M,s,tape,qhd,q_tl,table.
-legal_tape ls c rs →
-table = flatten ? (tuples_list (no_states M) (nhalt M) (graph_enum ?? (ntrans M))) →
-lift_tape ls c rs = low_tape_aux M tape →
-〈qhd,false〉::q_tl = m_bits_of_state (no_states M) (nhalt M) s →
-midtape STape (〈grid,false〉::ls)
- 〈qhd,false〉
- (q_tl@c::〈grid,false〉::table@〈grid,false〉::rs) =
- low_config M (mk_config ?? s tape).
-#ls #c #rs #M #s #tape #qhd #q_tl #table #Hlegal #Htable
-#Hlift #Hstate whd in match (low_config ??); <Hstate
-@eq_f3
- [@eq_f <(left_of_lift ls c rs) >Hlift //
- | cut (∀A.∀a,b:A.∀l1,l2. a::l1 = b::l2 → a=b)
- [#A #a #b #l1 #l2 #H destruct (H) %] #Hcut
- @(Hcut …Hstate)
- |@eq_f <(current_of_low … Hlegal Hlift) @eq_f @eq_f <Htable @eq_f @eq_f
- <(right_of_lift ls c rs Hlegal) >Hlift @right_of_low_tape
- ]
-qed.
-
-lemma unistep_true_to_low_step: ∀t1,t2.
- R_uni_step_true t1 t2 → low_step_R_true t1 t2.
-#t1 #t2 (* whd in ⊢ (%→%); *) #Huni_step * #n #posn #t #h * #qin #tape #eqt1
-cases (low_config_eq … eqt1)
-#low_left * #low_right * #table * #q_low_hd * #q_low_tl * #current_low
-***** #Hlow_left #Hlow_right #Htable #Hq_low #Hcurrent_low #Ht1
-letin trg ≝ (t 〈qin,current ? tape〉)
-letin qout_low ≝ (m_bits_of_state n h (\fst trg))
-letin qout_low_hd ≝ (hd ? qout_low 〈bit true,false〉)
-letin qout_low_tl ≝ (tail ? qout_low)
-letin low_act ≝ (low_action (\snd (t 〈qin,current ? tape〉)))
-letin low_cout ≝ (\fst low_act)
-letin low_m ≝ (\snd low_act)
-lapply (Huni_step n table q_low_hd (\fst qout_low_hd)
- current_low low_cout low_left low_right q_low_tl qout_low_tl low_m … Ht1)
- [@daemon
- |>Htable
- @(trans_to_match n h t 〈qin,current ? tape〉 … (refl …))
- >Hq_low >Hcurrent_low whd in match (mk_tuple ?????);
- >(eq_pair_fst_snd … (t …)) whd in ⊢ (??%?);
- >(eq_pair_fst_snd … (low_action …)) %
- |//
- |@daemon
- ]
--Ht1 #Huni_step lapply (Huni_step ? (refl …)) -Huni_step *
-#q_low_head_false * #ls1 * #rs1 * #c2 * *
-#Ht2 #Hlift #Hlegal %
- [whd in ⊢ (??%?); >q_low_head_false in Hq_low;
- whd in ⊢ ((???%)→?); generalize in match (h qin);
- #x #H destruct (H) %
- |>Ht2 whd in match (step FinBool ??);
- whd in match (trans ???);
- >(eq_pair_fst_snd … (t ?))
- @is_low_config // >Hlift
- <low_tape_move @eq_f2
- [>Hlow_left >Hlow_right >Hcurrent_low whd in ⊢ (??%%);
- cases (current …tape) [%] #b whd in ⊢ (??%%); %
- |whd in match low_cout; whd in match low_m; whd in match low_act;
- generalize in match (\snd (t ?)); * [%] * #b #mv
- whd in ⊢ (??(?(???%)?)%); cases mv %
- ]
- ]
-qed.
-
-definition low_step_R_false ≝ λt1,t2.
- ∀M:normalTM.
- ∀c: nconfig (no_states M).
- t1 = low_config M c → halt ? M (cstate … c) = true ∧ t1 = t2.
-
-lemma unistep_false_to_low_step: ∀t1,t2.
- R_uni_step_false t1 t2 → low_step_R_false t1 t2.
-#t1 #t2 (* whd in ⊢ (%→%); *) #Huni_step * #n #posn #t #h * #qin #tape #eqt1
-cases (low_config_eq … eqt1) #low_left * #low_right * #table * #q_low_hd * #q_low_tl * #current_low
-***** #_ #_ #_ #Hqin #_ #Ht1 whd in match (halt ???);
-cases (Huni_step (h qin) ?) [/2/] >Ht1 whd in ⊢ (??%?); @eq_f
-normalize in Hqin; destruct (Hqin) %
-qed.
-
-definition low_R ≝ λM,qstart,R,t1,t2.
- ∀tape1. t1 = low_config M (mk_config ?? qstart tape1) →
- ∃q,tape2.R tape1 tape2 ∧
- halt ? M q = true ∧ t2 = low_config M (mk_config ?? q tape2).
-
-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
-* #ta * #Hstar generalize in match q; -q
-@(star_ind_l ??????? Hstar)
- [#tb #q0 whd in ⊢ (%→?); #Htb #tape1 #Htb1
- cases (Htb … Htb1) -Htb #Hhalt #Htb
- <Htb >Htb1 @ex_intro
- [|%{tape1} %
- [ %
- [ whd @(ex_intro … 1) @(ex_intro … (mk_config … q0 tape1))
- % [|%] whd in ⊢ (??%?); >Hhalt %
- | @Hhalt ]
- | % ]
- ]
- |#tb #tc #td whd in ⊢ (%→?); #Htc #Hstar1 #IH
- #q #Htd #tape1 #Htb
- lapply (IH (\fst (trans ? M 〈q,current ? tape1〉)) Htd) -IH
- #IH cases (Htc … Htb); -Htc #Hhaltq
- whd in match (step FinBool ??); >(eq_pair_fst_snd ?? (trans ???))
- #Htc change with (mk_config ????) in Htc:(???(??%));
- cases (IH ? Htc) #q1 * #tape2 * * #HRTM #Hhaltq1 #Houtc
- @(ex_intro … q1) @(ex_intro … tape2) %
- [%
- [cases HRTM #k * #outc1 * #Hloop #Houtc1
- @(ex_intro … (S k)) @(ex_intro … outc1) %
- [>loopM_unfold >loop_S_false [2://] whd in match (step FinBool ??);
- >(eq_pair_fst_snd ?? (trans ???)) @Hloop
- |@Houtc1
- ]
- |@Hhaltq1]
- |@Houtc
- ]
- ]
-*)
-qed.
-
-theorem sem_universal2: ∀M:normalTM. ∀R.
- M ⊫ R → universalTM ⊫ (low_R M (start ? M) R).
-#M #R #HMR lapply (sem_universal … M (start ? M)) @WRealize_to_WRealize
-#t1 #t2 whd in ⊢ (%→%); #H #tape1 #Htape1 cases (H ? Htape1)
-#q * #tape2 * * #HRTM #Hhalt #Ht2 @(ex_intro … q) @(ex_intro … tape2)
-% [% [@(R_TM_to_R … HRTM) @HMR | //] | //]
-qed.
-
-axiom terminate_UTM: ∀M:normalTM.∀t.
- M ↓ t → universalTM ↓ (low_config M (mk_config ?? (start ? M) t)).
-