X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Fwmono.ma;h=d31ed62f551dabdb8e2963041c664f7b64e30415;hb=cd2f5b59215ea771ac137b9a7b115a05175f45d5;hp=6b7849709f870ebfc38d429c80ca03498fb075ae;hpb=b0378187bd0aeebb65502ad270264a980de4c8c0;p=helm.git diff --git a/matita/matita/lib/turing/wmono.ma b/matita/matita/lib/turing/wmono.ma index 6b7849709..d31ed62f5 100644 --- a/matita/matita/lib/turing/wmono.ma +++ b/matita/matita/lib/turing/wmono.ma @@ -10,6 +10,8 @@ V_____________________________________________________________*) include "basics/vectors.ma". +include "basics/finset.ma". +include "basics/core_notation/compose_2.ma". (* include "basics/relations.ma". *) record tape (sig:FinSet): Type[0] ≝ @@ -363,9 +365,12 @@ lemma eq_ctape_lift_conf_R : ∀sig,S1,S2,outc. #sig #S1 #S2 #outc cases outc #s #t % qed. +axiom daemon :∀P:Prop.P. + theorem sem_seq: ∀sig,M1,M2,R1,R2. Realize sig M1 R1 → Realize sig M2 R2 → Realize sig (seq sig M1 M2) (R1 ∘ R2). +@daemon (* this no longer works: TODO *) (* #sig #M1 #M2 #R1 #R2 #HR1 #HR2 #t #i #outc #Hloop cases (HR1 t) #k1 * #outc1 * #Hloop1 #HM1 cases (HR2 (ctape sig (states ? M1) outc1)) #k2 * #outc2 * #Hloop2 #HM2 @@ -395,5 +400,6 @@ cases (HR2 (ctape sig (states ? M1) outc1)) #k2 * #outc2 * #Hloop2 #HM2 | @(ex_intro … (ctape ? (FinSum (states ? M1) (states ? M2)) (lift_confL … outc1))) % // >eq_ctape_lift_conf_L >eq_ctape_lift_conf_R // ] +*) qed.