thys/turing_basic.thy
changeset 41 6d89ed67ba27
parent 40 a37a21f7ccf4
child 43 a8785fa80278
--- a/thys/turing_basic.thy	Sun Jan 13 11:29:33 2013 +0000
+++ b/thys/turing_basic.thy	Sun Jan 13 23:59:29 2013 +0000
@@ -31,6 +31,14 @@
 | "nth_of (x # xs) 0 = Some x"
 | "nth_of (x # xs) (Suc n) = nth_of xs n" 
 
+lemma nth_of_map [simp]:
+  shows "nth_of (map f p) n = (case (nth_of p n) of None \<Rightarrow> None | Some x \<Rightarrow> Some (f x))"
+apply(induct p arbitrary: n)
+apply(auto)
+apply(case_tac n)
+apply(auto)
+done
+
 fun 
   fetch :: "tprog \<Rightarrow> state \<Rightarrow> cell \<Rightarrow> instr"
 where
@@ -44,6 +52,14 @@
          Some i \<Rightarrow> i
        | None \<Rightarrow> (Nop, 0))"
 
+lemma fetch_Nil [simp]:
+  shows "fetch [] s b = (Nop, 0)"
+apply(case_tac s)
+apply(auto)
+apply(case_tac b)
+apply(auto)
+done
+
 fun 
   update :: "action \<Rightarrow> tape \<Rightarrow> tape"
 where 
@@ -80,6 +96,7 @@
 where
   "tm_wf p = (length p \<ge> 2 \<and> iseven (length p) \<and> (\<forall>(a, s) \<in> set p. s \<le> length p div 2))"
 
+
 (* FIXME: needed? *)
 lemma halt_lemma: 
   "\<lbrakk>wf LE; \<forall>n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
@@ -98,18 +115,18 @@
   "shift p n = (map (\<lambda> (a, s). (a, (if s = 0 then 0 else s + n))) p)"
 
 
-lemma [simp]: 
+lemma length_shift [simp]: 
   "length (shift p n) = length p"
 by (simp)
 
 fun 
   adjust :: "tprog \<Rightarrow> tprog"
 where
-  "adjust p = (map (\<lambda> (a, s). (a, if s = 0 then ((length p) div 2) + 1 else s)) p)"
+  "adjust p = map (\<lambda> (a, s). (a, if s = 0 then ((length p) div 2) + 1 else s)) p"
 
-lemma [simp]: 
+lemma length_adjust[simp]: 
   shows "length (adjust p) = length p"
-by (simp)
+by (induct p) (auto)
 
 definition
   tm_comp :: "tprog \<Rightarrow> tprog \<Rightarrow> tprog" ("_ |+| _" [0, 0] 100)
@@ -154,6 +171,12 @@
   where
   "assert_imp P Q = (\<forall>l r. P (l, r) \<longrightarrow> Q (l, r))"
 
+lemma holds_for_imp:
+  assumes "P holds_for c"
+  and "P \<mapsto> Q"
+  shows "Q holds_for c"
+using assms unfolding assert_imp_def by (case_tac c, auto)
+
 lemma test:
   assumes "is_final (steps (1, (l, r)) p n1)"
   and     "is_final (steps (1, (l, r)) p n2)"
@@ -182,6 +205,67 @@
     {P1} A |+| B {Q2}
 *}
 
+lemma before_final: 
+  assumes "steps (1, tp) A n = (0, tp')"
+  obtains n' where "\<not> is_final (steps (1, tp) A n')" and "steps (1, tp) A (Suc n') = (0, tp')"
+using assms 
+apply(induct n)
+apply(auto)
+by (metis is_final.simps step_red steps.simps steps_0 surj_pair)
+
+lemma t_merge_fetch_pre:
+  assumes "fetch A s b = (ac, ns)" "s \<le> length A div 2" "tm_wf A" "s \<noteq> 0" 
+  shows "fetch (adjustA |+| B) s b = (ac, if ns = 0 then Suc (length A div 2) else ns)"
+using assms
+unfolding tm_comp_def
+apply(induct A)
+apply(auto)
+apply(subgoal_tac "2 * (s - Suc 0) < length A \<and> Suc (2 * (s - Suc 0)) < length A")
+apply(auto simp: tm_wf_def iseven_def tm_comp_def split: if_splits cell.splits)
+oops
+
+lemma t_merge_pre_eq_step: 
+  "\<lbrakk>step (a, b, c) A = cf; tm_wf A; \<not> is_final cf\<rbrakk> 
+        \<Longrightarrow> step (a, b, c) (A |+| B) = cf"
+apply(subgoal_tac "a \<le> length A div 2 \<and> a \<noteq> 0")
+apply(simp)
+apply(case_tac "fetch A a (read c)", simp)
+apply(auto)
+oops
+
+lemma t_merge_pre_eq:  
+  "\<lbrakk>steps (Suc 0, tp) A stp = cf; \<not> is_final cf; tm_wf A\<rbrakk>
+  \<Longrightarrow> steps (Suc 0, tp) (A |+| B) stp = cf"
+apply(induct stp arbitrary: cf)
+apply(auto)[1]
+apply(auto)
+oops
+
+lemma t_merge_pre_halt_same: 
+  assumes a_ht: "steps (1, tp) A n = (0, tp')"
+  and a_wf: "t_correct A"
+  obtains n' where "steps (1, tp) (A |+| B) n' = (Suc (length A div 2), tp')"
+proof -
+  assume a: "\<And>n. steps (1, tp) (A |+| B) n = (Suc (length A div 2), tp') \<Longrightarrow> thesis"
+  
+  obtain stp' where "\<not> is_final (steps (1, tp) A stp')" and "steps (1, tp) A (Suc stp') = (0, tp')"
+  using a_ht before_final by blast
+  then have "steps (1, tp) (A |+| B) (Suc stp') = (Suc (length A div 2), tp')"
+    sorry (*using a_wf t_merge_pre_halt_same' by blast*)
+  with a show thesis by blast
+qed
+
+
+
+lemma steps_comp:
+  assumes a1: "steps (1, l, r) A n1 = (s1, l1, r1)"
+  and a2: "steps (1, l1, r1) B n2 = (s2, l2, r2)"
+  shows "steps (1, l, r) (A |+| B) (n1 + n2) = (s2, l2, r2)"
+using assms
+apply(induct n2)
+apply(simp)
+apply(simp add: tm_comp_def)
+oops
 
 lemma Hoare_plus: 
   assumes aimpb: "Q1 \<mapsto> P2"
@@ -191,83 +275,28 @@
   and B_halt : "{P2} B {Q2}"
   shows "{P1} A |+| B {Q2}"
 proof(rule HoareI)
-  fix a b
-  assume h: "P1 (a, b)"
-  then have "\<exists>n. is_final (steps (1, a, b) A n) \<and> Q1 holds_for (steps (1, a, b) A n)"
-    using A_halt unfolding Hoare_def by simp
-  then obtain n1 where "is_final (steps (1, a, b) A n1) \<and> Q1 holds_for (steps (1, a, b) A stp1)" ..
-  then show "\<exists>n. is_final (steps (1, a, b) (A |+| B) n) \<and> Q2 holds_for (steps (1, a, b) (A |+| B) n)"
-  proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE)
-    fix aa ba c
-    assume g1: "Q1 (ba, c)" 
-      and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)"
-    hence "P2 (ba, c)"
-      using aimpb apply(simp add: assert_imp_def)
-      done
-    hence "\<exists> stp. is_final (steps (Suc 0, ba, c) B stp) \<and> Q2 holds_for (steps (Suc 0, ba, c) B stp)"
-      using B_halt unfolding Hoare_def by simp
-    from this obtain stp2 where 
-      "is_final (steps (Suc 0, ba, c) B stp2) \<and> Q2 holds_for (steps (Suc 0, ba, c) B stp2)" ..
-    thus "?thesis"
-    proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE)
-      fix aa bb ca
-      assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)"
-      have "\<exists> stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)"
-        using g2 A_wf B_wf
-        sorry
-      moreover have "\<exists> stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)"
-        using g3 A_wf B_wf
-        sorry
-      ultimately show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'"
-        apply(erule_tac exE, erule_tac exE)
-        apply(rule_tac x = "stp + stpa" in exI, simp)
-        using g3 by simp
-    qed
-  qed
+  fix l r
+  assume h: "P1 (l, r)"
+  then obtain n1 where a: "is_final (steps (1, l, r) A n1)" and b: "Q1 holds_for (steps (1, l, r) A n1)"
+    using A_halt unfolding Hoare_def by auto
+  from b aimpb have "P2 holds_for (steps (1, l, r) A n1)"
+    by (rule holds_for_imp)
+  then obtain l' r' where "P2 (l', r')"
+    apply(auto)
+    apply(case_tac "steps (Suc 0, l, r) A n1")
+    apply(simp)
+    done
+  then obtain n2 where a: "is_final (steps (1, l', r') B n2)" and b: "Q2 holds_for (steps (1, l', r') B n2)"
+    using B_halt unfolding Hoare_def by auto
+  show "\<exists>n. is_final (steps (1, l, r) (A |+| B) n) \<and> Q2 holds_for (steps (1, l, r) (A |+| B) n)"
+    apply(rule_tac x="n1 + n2" in exI)
+    apply(rule conjI)
+    apply(simp)
+    apply(simp only: steps_add[symmetric])
+    sorry
 qed
 
-lemma Hoare_plus2: 
-  assumes A_wf : "tm_wf A"
-  and B_wf : "tm_wf B"
-  and A_halt : "{P1} A {Q1}"
-  and B_halt : "{P2} B {Q2}"
-  and aimpb: "Q1 \<mapsto> P2"
-  shows "{P1} A |+| B {Q2}"
-unfolding hoare_def
-proof(auto split: )
-  fix a b
-  assume h: "P1 (a, b)"
-  hence "\<exists>n. let (s, tp') = steps (Suc 0, a, b) A n in s = 0 \<and> Q1 tp'"
-    using A_halt unfolding hoare_def by simp
-  from this obtain stp1 where "let (s, tp') = steps (Suc 0, a, b) A stp1 in s = 0 \<and> Q1 tp'" ..
-  then show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'"
-  proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE)
-    fix aa ba c
-    assume g1: "Q1 (ba, c)" 
-      and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)"
-    hence "P2 (ba, c)"
-      using aimpb apply(simp add: assert_imp_def)
-      done
-    hence "\<exists> stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \<and> Q2 tp'"
-      using B_halt unfolding hoare_def by simp
-    from this obtain stp2 where "let (s, tp') = steps (Suc 0, ba, c) B stp2 in s = 0 \<and> Q2 tp'" ..
-    thus "?thesis"
-    proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE)
-      fix aa bb ca
-      assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)"
-      have "\<exists> stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)"
-        using g2 A_wf B_wf
-        sorry
-      moreover have "\<exists> stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)"
-        using g3 A_wf B_wf
-        sorry
-      ultimately show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'"
-        apply(erule_tac exE, erule_tac exE)
-        apply(rule_tac x = "stp + stpa" in exI, simp)
-        using g3 by simp
-    qed
-  qed
-qed
+