thys/Abacus_Hoare.thy
author Sebastiaan Joosten <sebastiaan.joosten@uibk.ac.at>
Fri, 21 Dec 2018 15:30:24 +0100
changeset 291 93db7414931d
parent 229 d8e6f0798e23
child 292 293e9c6f22e1
permissions -rwxr-xr-x
More naming of lemmas, cleanup of Abacus and NatBijection
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
229
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
theory Abacus_Hoare
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
imports Abacus
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
begin
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
type_synonym abc_assert = "nat list \<Rightarrow> bool"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
definition 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
  assert_imp :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" ("_ \<mapsto> _" [0, 0] 100)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
where
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
  "assert_imp P Q \<equiv> \<forall>xs. P xs \<longrightarrow> Q xs"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
fun abc_holds_for :: "(nat list \<Rightarrow> bool) \<Rightarrow> (nat \<times> nat list) \<Rightarrow> bool" ("_ abc'_holds'_for _" [100, 99] 100)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
where
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
  "P abc_holds_for (s, lm) = P lm"  
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
(* Hoare Rules *)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
(* halting case *)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
(*consts abc_Hoare_halt :: "abc_assert \<Rightarrow> abc_prog \<Rightarrow> abc_assert \<Rightarrow> bool" ("({(1_)}/ (_)/ {(1_)})" 50)*)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
fun abc_final :: "(nat \<times> nat list) \<Rightarrow> abc_prog \<Rightarrow> bool"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
  where 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
  "abc_final (s, lm) p = (s = length p)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
fun abc_notfinal :: "abc_conf \<Rightarrow> abc_prog \<Rightarrow> bool"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
  where
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
  "abc_notfinal (s, lm) p = (s < length p)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
definition 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
  abc_Hoare_halt :: "abc_assert \<Rightarrow> abc_prog \<Rightarrow> abc_assert \<Rightarrow> bool" ("({(1_)}/ (_)/ {(1_)})" 50)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
where
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
  "abc_Hoare_halt P p Q \<equiv> \<forall>lm. P lm \<longrightarrow> (\<exists>n. abc_final (abc_steps_l (0, lm) p n) p \<and> Q abc_holds_for (abc_steps_l (0, lm) p n))"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
lemma abc_Hoare_haltI:
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
  assumes "\<And>lm. P lm \<Longrightarrow> \<exists>n. abc_final (abc_steps_l (0, lm) p n) p \<and> Q abc_holds_for (abc_steps_l (0, lm) p n)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
  shows "{P} (p::abc_prog) {Q}"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
unfolding abc_Hoare_halt_def 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
using assms by auto
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
text {*
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
  {P} A {Q}   {Q} B {S} 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
  -----------------------------------------
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
  {P} A [+] B {S}
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
*}
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
definition
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
  abc_Hoare_unhalt :: "abc_assert \<Rightarrow> abc_prog \<Rightarrow> bool" ("({(1_)}/ (_)) \<up>" 50)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
where
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
  "abc_Hoare_unhalt P p \<equiv> \<forall>args. P args \<longrightarrow> (\<forall> n .abc_notfinal (abc_steps_l (0, args) p n) p)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
lemma abc_Hoare_unhaltI:
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
  assumes "\<And>args n. P args \<Longrightarrow> abc_notfinal (abc_steps_l (0, args) p n) p"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
  shows "{P} (p::abc_prog) \<up>"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
unfolding abc_Hoare_unhalt_def 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    54
using assms by auto
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
fun abc_inst_shift :: "abc_inst \<Rightarrow> nat \<Rightarrow> abc_inst"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
  where
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
  "abc_inst_shift (Inc m) n = Inc m" |
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
  "abc_inst_shift (Dec m e) n = Dec m (e + n)" |
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
  "abc_inst_shift (Goto m) n = Goto (m + n)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    61
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    62
fun abc_shift :: "abc_inst list \<Rightarrow> nat \<Rightarrow> abc_inst list" 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
  where
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
  "abc_shift xs n = map (\<lambda> x. abc_inst_shift x n) xs" 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
fun abc_comp :: "abc_inst list \<Rightarrow> abc_inst list \<Rightarrow> 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    67
                           abc_inst list" (infixl "[+]" 99)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    68
  where
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
  "abc_comp al bl = (let al_len = length al in 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    70
                           al @ abc_shift bl al_len)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    71
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
lemma abc_comp_first_step_eq_pre: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
  "s < length A
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
 \<Longrightarrow> abc_step_l (s, lm) (abc_fetch s (A [+] B)) = 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
    abc_step_l (s, lm) (abc_fetch s A)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
by(simp add: abc_step_l.simps abc_fetch.simps nth_append)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
lemma abc_before_final: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
  "\<lbrakk>abc_final (abc_steps_l (0, lm) p n) p; p \<noteq> []\<rbrakk>
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
  \<Longrightarrow> \<exists> n'. abc_notfinal (abc_steps_l (0, lm) p n') p \<and> 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
            abc_final (abc_steps_l (0, lm) p (Suc n')) p"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
proof(induct n)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
  case 0
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
  thus "?thesis"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
    by(simp add: abc_steps_l.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
next
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
  case (Suc n)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
  have ind: " \<lbrakk>abc_final (abc_steps_l (0, lm) p n) p; p \<noteq> []\<rbrakk> \<Longrightarrow> 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
    \<exists>n'. abc_notfinal (abc_steps_l (0, lm) p n') p \<and> abc_final (abc_steps_l (0, lm) p (Suc n')) p"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
    by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
  have final: "abc_final (abc_steps_l (0, lm) p (Suc n)) p" by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
  have notnull: "p \<noteq> []" by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    93
  show "?thesis"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    94
  proof(cases "abc_final (abc_steps_l (0, lm) p n) p")
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    95
    case True
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    96
    have "abc_final (abc_steps_l (0, lm) p n) p" by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    97
    then have "\<exists>n'. abc_notfinal (abc_steps_l (0, lm) p n') p \<and> abc_final (abc_steps_l (0, lm) p (Suc n')) p"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    98
      using ind notnull
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
      by simp
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
    thus "?thesis"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
      by simp
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   102
  next
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
    case False
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
    have "\<not> abc_final (abc_steps_l (0, lm) p n) p" by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
    from final this have "abc_notfinal (abc_steps_l (0, lm) p n) p" 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
      by(case_tac "abc_steps_l (0, lm) p n", simp add: abc_step_red2 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
        abc_step_l.simps abc_fetch.simps split: if_splits)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
    thus "?thesis"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   109
      using final
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
      by(rule_tac x = n in exI, simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
  qed
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
qed
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   113
    
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
lemma notfinal_Suc:
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
  "abc_notfinal (abc_steps_l (0, lm) A (Suc n)) A \<Longrightarrow>  
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
  abc_notfinal (abc_steps_l (0, lm) A n) A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
apply(case_tac "abc_steps_l (0, lm) A n")
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
apply(simp add: abc_step_red2 abc_fetch.simps abc_step_l.simps split: if_splits)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
done
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
lemma abc_comp_frist_steps_eq_pre: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   122
  assumes notfinal: "abc_notfinal (abc_steps_l (0, lm)  A n) A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
  and notnull: "A \<noteq> []"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   124
  shows "abc_steps_l (0, lm) (A [+] B) n = abc_steps_l (0, lm) A n"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
using notfinal
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
proof(induct n)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
  case 0
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   128
  thus "?case"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   129
    by(simp add: abc_steps_l.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
next
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
  case (Suc n)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   132
  have ind: "abc_notfinal (abc_steps_l (0, lm) A n) A \<Longrightarrow> abc_steps_l (0, lm) (A [+] B) n = abc_steps_l (0, lm) A n"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
    by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
  have h: "abc_notfinal (abc_steps_l (0, lm) A (Suc n)) A" by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   135
  then have a: "abc_notfinal (abc_steps_l (0, lm) A n) A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   136
    by(simp add: notfinal_Suc)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
  then have b: "abc_steps_l (0, lm) (A [+] B) n = abc_steps_l (0, lm) A n"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   138
    using ind by simp
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
  obtain s lm' where c: "abc_steps_l (0, lm) A n = (s, lm')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   140
    by (metis prod.exhaust)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
  then have d: "s < length A \<and> abc_steps_l (0, lm) (A [+] B) n = (s, lm')" 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
    using a b by simp
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
  thus "?case"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
    using c
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   145
    by(simp add: abc_step_red2 abc_fetch.simps abc_step_l.simps nth_append)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
qed
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
declare abc_shift.simps[simp del] abc_comp.simps[simp del]
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
lemma halt_steps2: "st \<ge> length A \<Longrightarrow> abc_steps_l (st, lm) A stp = (st, lm)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
apply(induct stp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
by(simp_all add: abc_step_red2 abc_steps_l.simps abc_step_l.simps abc_fetch.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   153
lemma halt_steps: "abc_steps_l (length A, lm) A n = (length A, lm)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   154
apply(induct n, simp add: abc_steps_l.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   155
apply(simp add: abc_step_red2 abc_step_l.simps nth_append abc_fetch.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
done
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
lemma abc_steps_add: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   159
  "abc_steps_l (as, lm) ap (m + n) = 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
         abc_steps_l (abc_steps_l (as, lm) ap m) ap n"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   161
apply(induct m arbitrary: n as lm, simp add: abc_steps_l.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
proof -
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   163
  fix m n as lm
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   164
  assume ind: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
    "\<And>n as lm. abc_steps_l (as, lm) ap (m + n) = 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
                   abc_steps_l (abc_steps_l (as, lm) ap m) ap n"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
  show "abc_steps_l (as, lm) ap (Suc m + n) = 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   168
             abc_steps_l (abc_steps_l (as, lm) ap (Suc m)) ap n"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   169
    apply(insert ind[of as lm "Suc n"], simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
    apply(insert ind[of as lm "Suc 0"], simp add: abc_steps_l.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
    apply(case_tac "(abc_steps_l (as, lm) ap m)", simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
    apply(simp add: abc_steps_l.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
    apply(case_tac "abc_step_l (a, b) (abc_fetch a ap)", 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
          simp add: abc_steps_l.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
    done
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
qed
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
lemma equal_when_halt: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
  assumes exc1: "abc_steps_l (s, lm) A na = (length A, lma)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
  and exc2: "abc_steps_l (s, lm) A nb = (length A, lmb)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
  shows "lma = lmb"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
proof(cases "na > nb")
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
  case True
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
  then obtain d where "na = nb + d"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   185
    by (metis add_Suc_right less_iff_Suc_add)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   186
  thus "?thesis" using assms halt_steps
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
    by(simp add: abc_steps_add)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
next
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
  case False
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
  then obtain d where "nb = na + d"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
    by (metis add.comm_neutral less_imp_add_positive nat_neq_iff)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
  thus "?thesis" using assms halt_steps
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
    by(simp add: abc_steps_add)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   194
qed 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
  
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   196
lemma abc_comp_frist_steps_halt_eq': 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   197
  assumes final: "abc_steps_l (0, lm) A n = (length A, lm')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   198
    and notnull: "A \<noteq> []"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
  shows "\<exists> n'. abc_steps_l (0, lm) (A [+] B) n' = (length A, lm')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
proof -
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
  have "\<exists> n'. abc_notfinal (abc_steps_l (0, lm) A n') A \<and> 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
    abc_final (abc_steps_l (0, lm) A (Suc n')) A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   203
    using assms
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
    by(rule_tac n = n in abc_before_final, simp_all)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
  then obtain na where a:
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
    "abc_notfinal (abc_steps_l (0, lm) A na) A \<and> 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
            abc_final (abc_steps_l (0, lm) A (Suc na)) A" ..
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   208
  obtain sa lma where b: "abc_steps_l (0, lm) A na = (sa, lma)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
    by (metis prod.exhaust)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   210
  then have c: "abc_steps_l (0, lm) (A [+] B) na = (sa, lma)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
    using a abc_comp_frist_steps_eq_pre[of lm A na B] assms 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   212
    by simp
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
  have d: "sa < length A" using b a by simp
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
  then have e: "abc_step_l (sa, lma) (abc_fetch sa (A [+] B)) = 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
    abc_step_l (sa, lma) (abc_fetch sa A)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
    by(rule_tac abc_comp_first_step_eq_pre)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   217
  from a have "abc_steps_l (0, lm) A (Suc na) = (length A, lm')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
    using final equal_when_halt
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
    by(case_tac "abc_steps_l (0, lm) A (Suc na)" , simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
  then have "abc_steps_l (0, lm) (A [+] B) (Suc na) = (length A, lm')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
    using a b c e
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   222
    by(simp add: abc_step_red2)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
  thus "?thesis"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
    by blast
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
qed
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   226
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   227
lemma abc_exec_null: "abc_steps_l sam [] n = sam"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
apply(cases sam)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
apply(induct n)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   230
apply(auto simp: abc_step_red2)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
apply(auto simp: abc_step_l.simps abc_steps_l.simps abc_fetch.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   232
done
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   234
lemma abc_comp_frist_steps_halt_eq: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
  assumes final: "abc_steps_l (0, lm) A n = (length A, lm')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
  shows "\<exists> n'. abc_steps_l (0, lm) (A [+] B) n' = (length A, lm')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
using final
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
apply(case_tac "A = []")
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   239
apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps abc_exec_null)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   240
apply(rule_tac abc_comp_frist_steps_halt_eq', simp_all)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
done
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
291
93db7414931d More naming of lemmas, cleanup of Abacus and NatBijection
Sebastiaan Joosten <sebastiaan.joosten@uibk.ac.at>
parents: 229
diff changeset
   243
229
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
lemma abc_comp_second_step_eq: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   245
  assumes exec: "abc_step_l (s, lm) (abc_fetch s B) = (sa, lma)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   246
  shows "abc_step_l (s + length A, lm) (abc_fetch (s + length A) (A [+] B))
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
         = (sa + length A, lma)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   248
using assms
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   249
apply(auto simp: abc_step_l.simps abc_fetch.simps nth_append abc_comp.simps abc_shift.simps split : if_splits )
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   250
apply(case_tac [!] "B ! s", auto simp: Let_def)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   251
done
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   253
lemma abc_comp_second_steps_eq:
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   254
  assumes exec: "abc_steps_l (0, lm) B n = (sa, lm')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   255
  shows "abc_steps_l (length A, lm) (A [+] B) n = (sa + length A, lm')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   256
using assms
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   257
proof(induct n arbitrary: sa lm')
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   258
  case 0
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
  thus "?case"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   260
    by(simp add: abc_steps_l.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
next
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
  case (Suc n)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   263
  have ind: "\<And>sa lm'. abc_steps_l (0, lm) B n = (sa, lm') \<Longrightarrow> 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
    abc_steps_l (length A, lm) (A [+] B) n = (sa + length A, lm')" by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   265
  have exec: "abc_steps_l (0, lm) B (Suc n) = (sa, lm')" by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   266
  obtain sb lmb where a: " abc_steps_l (0, lm) B n = (sb, lmb)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   267
    by (metis prod.exhaust)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
 then have "abc_steps_l (length A, lm) (A [+] B) n = (sb + length A, lmb)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
   using ind by simp
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   270
 moreover have "abc_step_l (sb + length A, lmb) (abc_fetch (sb + length A) (A [+] B)) = (sa + length A, lm') "
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
   using a exec abc_comp_second_step_eq
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   272
   by(simp add: abc_step_red2)    
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   273
 ultimately show "?case"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   274
   by(simp add: abc_step_red2)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   275
qed
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   276
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   277
lemma length_abc_comp[simp, intro]: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
  "length (A [+] B) = length A + length B"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   279
by(auto simp: abc_comp.simps abc_shift.simps)   
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   280
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   281
lemma abc_Hoare_plus_halt : 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
  assumes A_halt : "{P} (A::abc_prog) {Q}"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   283
  and B_halt : "{Q} (B::abc_prog) {S}"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   284
  shows "{P} (A [+] B) {S}"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   285
proof(rule_tac abc_Hoare_haltI)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   286
  fix lm
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   287
  assume a: "P lm"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   288
  then obtain na lma where 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   289
    "abc_final (abc_steps_l (0, lm) A na) A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   290
    and b: "abc_steps_l (0, lm) A na = (length A, lma)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   291
    and c: "Q abc_holds_for (length A, lma)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   292
    using A_halt unfolding abc_Hoare_halt_def
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   293
    by (metis (full_types) abc_final.simps abc_holds_for.simps prod.exhaust)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   294
  have "\<exists> n. abc_steps_l (0, lm) (A [+] B) n = (length A, lma)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   295
    using abc_comp_frist_steps_halt_eq b
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   296
    by(simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   297
  then obtain nx where h1: "abc_steps_l (0, lm) (A [+] B) nx = (length A, lma)" ..   
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   298
  from c have "Q lma"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   299
    using c unfolding abc_holds_for.simps
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   300
    by simp
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   301
  then obtain nb lmb where
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   302
    "abc_final (abc_steps_l (0, lma) B nb) B"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   303
    and d: "abc_steps_l (0, lma) B nb = (length B, lmb)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   304
    and e: "S abc_holds_for (length B, lmb)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   305
    using B_halt unfolding abc_Hoare_halt_def
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   306
    by (metis (full_types) abc_final.simps abc_holds_for.simps prod.exhaust)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   307
  have h2: "abc_steps_l (length A, lma) (A [+] B) nb = (length B + length A, lmb)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   308
    using d abc_comp_second_steps_eq
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   309
    by simp
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   310
  thus "\<exists>n. abc_final (abc_steps_l (0, lm) (A [+] B) n) (A [+] B) \<and>
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   311
    S abc_holds_for abc_steps_l (0, lm) (A [+] B) n"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   312
    using h1 e
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   313
    by(rule_tac x = "nx + nb" in exI, simp add: abc_steps_add)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   314
qed
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   315
 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   316
lemma abc_unhalt_append_eq:
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   317
  assumes unhalt: "{P} (A::abc_prog) \<up>"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
  and P: "P args"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   319
  shows "abc_steps_l (0, args) (A [+] B) stp = abc_steps_l (0, args) A stp"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
proof(induct stp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   321
  case 0
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   322
  thus "?case"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   323
    by(simp add: abc_steps_l.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   324
next
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   325
  case (Suc stp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   326
  have ind: "abc_steps_l (0, args) (A [+] B) stp = abc_steps_l (0, args) A stp"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   327
    by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
  obtain s nl where a: "abc_steps_l (0, args) A stp = (s, nl)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   329
    by (metis prod.exhaust)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   330
  then have b: "s < length A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
    using unhalt P
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
    apply(auto simp: abc_Hoare_unhalt_def)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   333
    by (metis abc_notfinal.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
  thus "?case"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   335
    using a ind
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   336
    by(simp add: abc_step_red2 abc_step_l.simps abc_fetch.simps nth_append abc_comp.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
qed
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
lemma abc_Hoare_plus_unhalt1: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
  "{P} (A::abc_prog) \<up> \<Longrightarrow> {P} (A [+] B) \<up>"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
apply(rule_tac abc_Hoare_unhaltI)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   342
apply(frule_tac args = args and B = B and stp = n in abc_unhalt_append_eq)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
apply(simp_all add: abc_Hoare_unhalt_def)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   344
apply(erule_tac x = args in allE, simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   345
apply(erule_tac x = n in allE)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
apply(case_tac "(abc_steps_l (0, args) A n)", simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
done
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   349
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   350
lemma notfinal_all_before:
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   351
  "\<lbrakk>abc_notfinal (abc_steps_l (0, args) A x) A; y\<le>x \<rbrakk>
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   352
  \<Longrightarrow> abc_notfinal (abc_steps_l (0, args) A y) A "
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
apply(subgoal_tac "\<exists> d. x = y + d", auto)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
apply(case_tac "abc_steps_l (0, args) A y",simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
apply(rule_tac classical, simp add: abc_steps_add leI halt_steps2)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   356
by arith
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   357
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   358
lemma abc_Hoare_plus_unhalt2':
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   359
  assumes unhalt: "{Q} (B::abc_prog) \<up>"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   360
   and halt: "{P} (A::abc_prog) {Q}"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   361
   and notnull: "A \<noteq> []"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   362
   and P: "P args" 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   363
   shows "abc_notfinal (abc_steps_l (0, args) (A [+] B) n) (A [+] B)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   364
proof -
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   365
  obtain st nl stp where a: "abc_final (abc_steps_l (0, args) A stp) A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   366
    and b: "Q abc_holds_for (length A, nl)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   367
    and c: "abc_steps_l (0, args) A stp = (st, nl)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   368
    using halt P unfolding abc_Hoare_halt_def
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   369
    by (metis abc_holds_for.simps prod.exhaust)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   370
  thm abc_before_final
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   371
  obtain stpa where d: 
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   372
    "abc_notfinal (abc_steps_l (0, args) A stpa) A \<and> abc_final (abc_steps_l (0, args) A (Suc stpa)) A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   373
    using a notnull abc_before_final[of args A stp]
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   374
    by(auto)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   375
  thus "?thesis"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   376
  proof(cases "n < Suc stpa")
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   377
    case True
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   378
    have h: "n < Suc stpa" by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
    then have "abc_notfinal (abc_steps_l (0, args) A n) A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   380
      using d
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   381
      by(rule_tac notfinal_all_before, auto)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   382
    moreover then have "abc_steps_l (0, args) (A [+] B) n = abc_steps_l (0, args) A n"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   383
      using notnull
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   384
      by(rule_tac abc_comp_frist_steps_eq_pre, simp_all)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   385
    ultimately show "?thesis"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   386
      by(case_tac "abc_steps_l (0, args) A n", simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   387
  next
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   388
    case False
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   389
    have "\<not> n < Suc stpa" by fact
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   390
    then obtain d where i1: "n = Suc stpa + d"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   391
      by (metis add_Suc less_iff_Suc_add not_less_eq)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   392
    have "abc_steps_l (0, args) A (Suc stpa) = (length A, nl)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   393
      using d a c
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   394
      apply(case_tac "abc_steps_l (0, args) A stp", simp add: equal_when_halt)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   395
      by(case_tac "abc_steps_l (0, args) A (Suc stpa)", simp add: equal_when_halt)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   396
    moreover have  "abc_steps_l (0, args) (A [+] B) stpa = abc_steps_l (0, args) A stpa"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   397
      using notnull d
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   398
      by(rule_tac abc_comp_frist_steps_eq_pre, simp_all)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   399
    ultimately have i2: "abc_steps_l (0, args) (A [+] B) (Suc stpa) = (length A, nl)"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   400
      using d
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   401
      apply(case_tac "abc_steps_l (0, args) A stpa", simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   402
      by(simp add: abc_step_red2 abc_steps_l.simps abc_fetch.simps abc_comp.simps nth_append)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   403
    obtain s' nl' where i3:"abc_steps_l (0, nl) B d = (s', nl')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   404
      by (metis prod.exhaust)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   405
    then have i4: "abc_steps_l (0, args) (A [+] B) (Suc stpa + d) = (length A + s', nl')"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   406
      using i2  apply(simp only: abc_steps_add)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   407
      using abc_comp_second_steps_eq[of nl B d s' nl']
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   408
      by simp
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   409
    moreover have "s' < length B"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   410
      using unhalt b i3
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   411
      apply(simp add: abc_Hoare_unhalt_def)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   412
      apply(erule_tac x = nl in allE, simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   413
      by(erule_tac x = d in allE, simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   414
    ultimately show "?thesis"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   415
      using i1
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   416
      by(simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   417
  qed
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   418
qed
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   419
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   420
lemma abc_comp_null_left[simp]: "[] [+] A = A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   421
apply(induct A)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   422
apply(case_tac [2] a)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   423
apply(auto simp: abc_comp.simps abc_shift.simps abc_inst_shift.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   424
done
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   425
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   426
lemma abc_comp_null_right[simp]: "A [+] [] = A"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   427
apply(induct A)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   428
apply(case_tac [2] a)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   429
apply(auto simp: abc_comp.simps abc_shift.simps abc_inst_shift.simps)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   430
done
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   431
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   432
lemma abc_Hoare_plus_unhalt2:
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   433
  "\<lbrakk>{Q} (B::abc_prog)\<up>; {P} (A::abc_prog) {Q}\<rbrakk>\<Longrightarrow> {P} (A [+] B) \<up>"
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   434
apply(case_tac "A = []")
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   435
apply(simp add: abc_Hoare_halt_def abc_Hoare_unhalt_def abc_exec_null)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   436
apply(rule_tac abc_Hoare_unhaltI)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   437
apply(erule_tac abc_Hoare_plus_unhalt2', simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   438
apply(simp, simp)
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   439
done
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   440
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   441
end
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
d8e6f0798e23 much simplified version of Recursive.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443