thys/Hoare_tm3.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Sat, 13 Sep 2014 04:39:07 +0100
changeset 26 1cde7bf45858
parent 20 e04123f4bacc
permissions -rw-r--r--
deleted *~ files
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
header {* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
  Separation logic for Turing Machines
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
theory Hoare_tm3
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
imports Hoare_gen My_block Data_slot FMap
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
begin
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
notation FMap.lookup (infixl "$" 999)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
fun pretty_terms ctxt trms =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
  Pretty.block (Pretty.commas (map (Syntax.pretty_term ctxt) trms))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
structure StepRules = Named_Thms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
  (val name = @{binding "step"}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
   val description = "Theorems for hoare rules for every step")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
structure FwdRules = Named_Thms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
  (val name = @{binding "fwd"}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
   val description = "Theorems for fwd derivation of seperation implication")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
setup {* StepRules.setup *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
setup {* FwdRules.setup *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
method_setup prune = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
  {* Scan.succeed (SIMPLE_METHOD' o (K (K prune_params_tac))) *} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
  "pruning parameters"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
lemma int_add_C: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
  "x + (y::int) = y + x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
lemma int_add_A : "(x + y) + z = x + (y + (z::int))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
lemma int_add_LC: "x + (y + (z::int)) = y + (x + z)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
lemmas int_add_ac = int_add_A int_add_C int_add_LC
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
section {* Operational Semantics of TM *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
    50
datatype action = W0 | W1 | L | R
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
    51
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
    52
type_synonym state = nat
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
    53
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
    54
(* FIXME: I think Block should be changed to cell *)
19
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
datatype Block = Oc | Bk
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
(* the successor state when tape symbol is Bk or Oc, respectively *)
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
    58
type_synonym tm_inst = "(action \<times> state) \<times> (action \<times> state)"
19
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
(* - number of faults (nat)
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
    61
   - TM program (nat f\<rightharpoonup> tm_inst)
19
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    62
   - current state (nat)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
   - position of head (int)
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
    64
   - tape (int f\<rightharpoonup> Block)
19
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
*)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
type_synonym tconf = "nat \<times> (nat f\<rightharpoonup> tm_inst) \<times> nat \<times> int \<times> (int f\<rightharpoonup> Block)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    67
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    68
(* updates the position/tape according to an action *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
fun 
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
    70
  next_tape :: "action \<Rightarrow> (int \<times>  (int f\<rightharpoonup> Block)) \<Rightarrow> (int \<times>  (int f\<rightharpoonup> Block))"
19
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    71
where 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
  "next_tape W0 (pos, m) = (pos, m(pos f\<mapsto> Bk))" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
  "next_tape W1 (pos, m) = (pos, m(pos f\<mapsto> Oc))" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
  "next_tape L  (pos, m) = (pos - 1, m)" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
  "next_tape R  (pos, m) = (pos + 1, m)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
fun tstep :: "tconf \<Rightarrow> tconf"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
  where "tstep (faults, prog, pc, pos, m) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
              (case (prog $ pc) of
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
                  Some ((action0, pc0), (action1, pc1)) \<Rightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
                     case (m $ pos) of
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
                       Some Bk \<Rightarrow> (faults, prog, pc0, next_tape action0 (pos, m)) |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
                       Some Oc \<Rightarrow> (faults, prog, pc1, next_tape action1 (pos, m)) |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
                       None \<Rightarrow> (faults + 1, prog, pc, pos, m)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
                | None \<Rightarrow> (faults + 1, prog, pc, pos, m))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
lemma tstep_splits: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
  "(P (tstep s)) = ((\<forall> faults prog pc pos m action0 pc0 action1 pc1. 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
                          s = (faults, prog, pc, pos, m) \<longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
                          prog $ pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
                          m $ pos = Some Bk \<longrightarrow> P (faults, prog, pc0, next_tape action0 (pos, m))) \<and>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
                    (\<forall> faults prog pc pos m action0 pc0 action1 pc1. 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    93
                          s = (faults, prog, pc, pos, m) \<longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    94
                          prog $ pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    95
                          m $ pos = Some Oc \<longrightarrow> P (faults, prog, pc1, next_tape action1 (pos, m))) \<and>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    96
                    (\<forall> faults prog pc pos m action0 pc0 action1 pc1. 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    97
                          s = (faults, prog, pc, pos, m) \<longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    98
                          prog $ pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
                          m $ pos = None \<longrightarrow> P (faults + 1, prog, pc, pos, m)) \<and>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
                    (\<forall> faults prog pc pos m . 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
                          s =  (faults, prog, pc, pos, m) \<longrightarrow>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   102
                          prog $ pc  = None \<longrightarrow> P (faults + 1, prog, pc, pos, m))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
                   )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
  by (cases s) (auto split: option.splits Block.splits)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
datatype tresource = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
    TM int Block      (* at the position int of the tape is a Bl or Oc *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
  | TC nat tm_inst    (* at the address nat is a certain instruction *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   109
  | TAt nat           (* TM is at state nat*)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
  | TPos int          (* head of TM is at position int *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
  | TFaults nat       (* there are nat faults *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   113
definition "tprog_set prog = {TC i inst | i inst. prog $ i = Some inst}"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
definition "tpc_set pc = {TAt pc}"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
definition "tmem_set m = {TM i n | i n. m $ i = Some n}"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
definition "tpos_set pos = {TPos pos}"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
definition "tfaults_set faults = {TFaults faults}"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
lemmas tpn_set_def = tprog_set_def tpc_set_def tmem_set_def tfaults_set_def tpos_set_def
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
fun trset_of :: "tconf \<Rightarrow> tresource set"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   122
  where "trset_of (faults, prog, pc, pos, m) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
               tprog_set prog \<union> tpc_set pc \<union> tpos_set pos \<union> tmem_set m \<union> tfaults_set faults"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   124
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
interpretation tm: sep_exec tstep trset_of .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   128
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   129
section {* Assembly language for TMs and its program logic *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
subsection {* The assembly language *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   132
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
datatype tpg = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
   TInstr tm_inst
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   135
 | TLabel nat
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   136
 | TSeq tpg tpg
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
 | TLocal "nat \<Rightarrow> tpg"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   138
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
notation TLocal (binder "TL " 10)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   140
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
abbreviation 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
  tprog_instr :: "tm_inst \<Rightarrow> tpg" ("\<guillemotright> _" [61] 61)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
where "\<guillemotright> i \<equiv> TInstr i"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   145
abbreviation tprog_seq :: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
  "tpg \<Rightarrow> tpg \<Rightarrow> tpg" (infixr ";" 52)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
where "c1 ; c2 \<equiv> TSeq c1 c2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
definition "sg e = (\<lambda>s. s = e)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
type_synonym tassert = "tresource set \<Rightarrow> bool"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   153
abbreviation 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   154
  t_hoare :: "tassert \<Rightarrow> tassert  \<Rightarrow> tassert \<Rightarrow> bool" ("(\<lbrace>(1_)\<rbrace>/ (_)/ \<lbrace>(1_)\<rbrace>)" 50)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   155
where 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
  "\<lbrace>p\<rbrace> c \<lbrace>q\<rbrace> == sep_exec.Hoare_gen tstep trset_of p c q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
abbreviation it_hoare ::
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   159
  "('a::sep_algebra \<Rightarrow> tresource set \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> (tresource set \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
      ("(1_).(\<lbrace>(1_)\<rbrace>/ (_)/ \<lbrace>(1_)\<rbrace>)" 50)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   161
where "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace> == sep_exec.IHoare tstep trset_of I P c Q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   163
(*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   164
primrec tpg_len :: "tpg \<Rightarrow> nat" where 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
  "tpg_len (TInstr ai) = 1" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
  "tpg_len (TSeq p1 p2) = tpg_len p1 + tpg_len " |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
  "tpg_len (TLocal fp) = tpg_len (fp 0)" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   168
  "tpg_len (TLabel l) = 0" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   169
*)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
primrec tassemble_to :: "tpg \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> tassert" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
  where 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
  "tassemble_to (TInstr ai) i j = (sg ({TC i ai}) ** <(j = i + 1)>)" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
  "tassemble_to (TSeq p1 p2) i j = (EXS j'. (tassemble_to p1 i j') ** (tassemble_to p2 j' j))" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
  "tassemble_to (TLocal fp) i j  = (EXS l. (tassemble_to (fp l) i j))" | 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
  "tassemble_to (TLabel l) i j = <(i = j \<and> j = l)>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
declare tassemble_to.simps [simp del]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
lemmas tasmp = tassemble_to.simps (2, 3, 4)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
abbreviation 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
  tasmb_to :: "nat \<Rightarrow> tpg \<Rightarrow> nat \<Rightarrow> tassert" ("_ :[ _ ]: _" [60, 60, 60] 60)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
  where 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   185
  "i :[ tpg ]: j \<equiv> tassemble_to tpg i j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   186
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
lemma EXS_intro: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
  assumes h: "(P x) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
  shows "(EXS x. P(x)) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
  by (smt h pred_ex_def sep_conj_impl)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
lemma EXS_elim: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
  assumes "(EXS x. P x) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   194
  obtains x where "P x s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
  by (metis assms pred_ex_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   196
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   197
lemma EXS_eq:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   198
  fixes x
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
  assumes h: "Q (p x)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
  and h1: "\<And> y s. \<lbrakk>p y s\<rbrakk> \<Longrightarrow> y = x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
  shows "p x = (EXS x. p x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
proof
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   203
  fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
  show "p x s = (EXS x. p x) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
  proof
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
    assume "p x s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
    thus "(EXS x. p x) s" by (auto simp:pred_ex_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   208
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
    assume "(EXS x. p x) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   210
    thus "p x s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
    proof(rule EXS_elim)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   212
      fix y
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
      assume "p y s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
      from this[unfolded h1[OF this]] show "p x s" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   217
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
definition "tpg_fold tpgs = foldr TSeq (butlast tpgs) (last tpgs)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
lemma tpg_fold_sg: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   222
  "tpg_fold [tpg] = tpg"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
  by (simp add:tpg_fold_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
lemma tpg_fold_cons: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   226
  assumes h: "tpgs \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   227
  shows "tpg_fold (tpg#tpgs) = (tpg; (tpg_fold tpgs))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
  using h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
proof(induct tpgs arbitrary:tpg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   230
  case (Cons tpg1 tpgs1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
  thus ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   232
  proof(cases "tpgs1 = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   234
    thus ?thesis by (simp add:tpg_fold_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
    proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   239
      have eq_1: "butlast (tpg # tpg1 # tpgs1) = tpg # (butlast (tpg1 # tpgs1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   240
        by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
      from False have eq_2: "last (tpg # tpg1 # tpgs1) = last (tpg1 # tpgs1)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
        by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
      have eq_3: "foldr (op ;) (tpg # butlast (tpg1 # tpgs1)) (last (tpg1 # tpgs1)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
            (tpg ; (foldr (op ;) (butlast (tpg1 # tpgs1)) (last (tpg1 # tpgs1))))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   245
        by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   246
      show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
        apply (subst (1) tpg_fold_def, unfold eq_1 eq_2 eq_3)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   248
        by (fold tpg_fold_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   249
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   250
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   251
qed auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   253
lemmas tpg_fold_simps = tpg_fold_sg tpg_fold_cons
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   254
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   255
lemma tpg_fold_app:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   256
  assumes h1: "tpgs1 \<noteq> []" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   257
  and h2: "tpgs2 \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   258
  shows "i:[(tpg_fold (tpgs1 @ tpgs2))]:j = i:[(tpg_fold (tpgs1); tpg_fold tpgs2)]:j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
  using h1 h2
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   260
proof(induct tpgs1 arbitrary: i j tpgs2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
  case (Cons tpg1' tpgs1')
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
  thus ?case (is "?L = ?R")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   263
  proof(cases "tpgs1' = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   265
    from h2 have "(tpgs1' @ tpgs2) \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   266
      by (metis Cons.prems(2) Nil_is_append_conv) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   267
    have "?L = (i :[ tpg_fold (tpg1' # (tpgs1' @ tpgs2)) ]: j )" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
    also have "\<dots> =  (i:[(tpg1'; (tpg_fold (tpgs1' @ tpgs2)))]:j )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
      by (simp add:tpg_fold_cons[OF `(tpgs1' @ tpgs2) \<noteq> []`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   270
    also have "\<dots> = ?R"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
    proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   272
      have "(EXS j'. i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   273
              (EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \<and>* j'a :[ tpg_fold tpgs1' ]: j') \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   274
                               j' :[ tpg_fold tpgs2 ]: j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   275
      proof(default+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   276
        fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   277
        assume "(EXS j'. i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
        thus "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \<and>* j'a :[ tpg_fold tpgs1' ]: j') \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   279
                  j' :[ tpg_fold tpgs2 ]: j) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   280
        proof(elim EXS_elim)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   281
          fix j'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
          assume "(i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   283
          from this[unfolded Cons(1)[OF False Cons(3)] tassemble_to.simps]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   284
          show "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \<and>* j'a :[ tpg_fold tpgs1' ]: j') \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   285
                           j' :[ tpg_fold tpgs2 ]: j) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   286
          proof(elim sep_conjE EXS_elim)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   287
            fix x y j'a xa ya
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   288
            assume h: "(i :[ tpg1' ]: j') x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   289
                      "x ## y" "s = x + y"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   290
                      "(j' :[ tpg_fold tpgs1' ]: j'a) xa"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   291
                      "(j'a :[ tpg_fold tpgs2 ]: j) ya"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   292
                      " xa ## ya" "y = xa + ya"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   293
            show "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   294
                                j'a :[ tpg_fold tpgs1' ]: j') \<and>* j' :[ tpg_fold tpgs2 ]: j) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   295
               (is "(EXS j'. (?P j' \<and>* ?Q j')) s")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   296
            proof(rule EXS_intro[where x = "j'a"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   297
              from `(j'a :[ tpg_fold tpgs2 ]: j) ya` have "(?Q j'a) ya" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   298
              moreover have "(?P j'a) (x + xa)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   299
              proof(rule EXS_intro[where x = j'])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   300
                have "x + xa = x + xa" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   301
                moreover from `x ## y` `y = xa + ya` `xa ## ya` 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   302
                have "x ## xa" by (metis sep_disj_addD) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   303
                moreover note `(i :[ tpg1' ]: j') x` `(j' :[ tpg_fold tpgs1' ]: j'a) xa`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   304
                ultimately show "(i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold tpgs1' ]: j'a) (x + xa)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   305
                  by (auto intro!:sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   306
              qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   307
              moreover from `x##y` `y = xa + ya` `xa ## ya` 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   308
              have "(x + xa) ## ya"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   309
                by (metis sep_disj_addI1 sep_disj_commuteI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   310
              moreover from `s = x + y` `y = xa + ya`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   311
              have "s = (x + xa) + ya"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   312
                by (metis h(2) h(6) sep_add_assoc sep_disj_addD1 sep_disj_addD2) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   313
              ultimately show "(?P j'a \<and>* ?Q j'a) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   314
                by (auto intro!:sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   315
            qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   316
          qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   317
        qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
      next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   319
        fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
        assume "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \<and>* j'a :[ tpg_fold tpgs1' ]: j') \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   321
                                    j' :[ tpg_fold tpgs2 ]: j) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   322
        thus "(EXS j'. i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   323
        proof(elim EXS_elim sep_conjE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   324
          fix j' x y j'a xa ya
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   325
          assume h: "(j' :[ tpg_fold tpgs2 ]: j) y"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   326
                    "x ## y" "s = x + y" "(i :[ tpg1' ]: j'a) xa"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   327
                    "(j'a :[ tpg_fold tpgs1' ]: j') ya" "xa ## ya" "x = xa + ya"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
          show "(EXS j'. i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   329
          proof(rule EXS_intro[where x = j'a])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   330
            from `(i :[ tpg1' ]: j'a) xa` have "(i :[ tpg1' ]: j'a) xa" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
            moreover have "(j'a :[ tpg_fold (tpgs1' @ tpgs2) ]: j) (ya + y)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
            proof(unfold Cons(1)[OF False Cons(3)] tassemble_to.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   333
              show "(EXS j'. j'a :[ tpg_fold tpgs1' ]: j' \<and>* j' :[ tpg_fold tpgs2 ]: j) (ya + y)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
              proof(rule EXS_intro[where x = "j'"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   335
                from `x ## y` `x = xa + ya` `xa ## ya`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   336
                have "ya ## y" by (metis sep_add_disjD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
                moreover have "ya + y = ya + y" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
                moreover note `(j'a :[ tpg_fold tpgs1' ]: j') ya` 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
                               `(j' :[ tpg_fold tpgs2 ]: j) y`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
                ultimately show "(j'a :[ tpg_fold tpgs1' ]: j' \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
                                 j' :[ tpg_fold tpgs2 ]: j) (ya + y)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   342
                  by (auto intro!:sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
              qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   344
            qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   345
            moreover from `s = x + y` `x = xa + ya`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
            have "s = xa + (ya + y)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
              by (metis h(2) h(6) sep_add_assoc sep_add_disjD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
            moreover from `xa ## ya` `x ## y` `x = xa + ya`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   349
            have "xa ## (ya + y)" by (metis sep_disj_addI3) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   350
            ultimately show "(i :[ tpg1' ]: j'a \<and>* j'a :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   351
              by (auto intro!:sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   352
          qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
        qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
      thus ?thesis 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   356
        by (simp add:tassemble_to.simps, unfold tpg_fold_cons[OF False], 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   357
             unfold tassemble_to.simps, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   358
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   359
    finally show ?thesis . 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   360
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   361
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   362
    thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   363
      by (simp add:tpg_fold_cons[OF Cons(3)] tpg_fold_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   364
  qed 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   365
qed auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   366
 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   367
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   368
subsection {* Assertions and program logic for this assembly language *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   369
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   370
definition "st l = sg (tpc_set l)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   371
definition "ps p = sg (tpos_set p)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   372
definition "tm a v = sg ({TM a v})"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   373
definition "one i = tm i Oc"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   374
definition "zero i= tm i Bk"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   375
definition "any i = (EXS v. tm i v)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   376
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   377
declare trset_of.simps[simp del]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   378
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
lemma stimes_sgD: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   380
  "(sg x \<and>* q) s \<Longrightarrow> q (s - x) \<and> x \<subseteq> s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   381
  apply(erule_tac sep_conjE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   382
  apply(unfold set_ins_def sg_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   383
  by (metis Diff_Int Diff_cancel Diff_empty Un_Diff sup.cobounded1 sup_bot.left_neutral sup_commute)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   384
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   385
lemma stD: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   386
  "(st i \<and>* r) (trset_of (ft, prog, i', pos, mem)) \<Longrightarrow> i' = i"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   387
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   388
  assume "(st i \<and>* r) (trset_of (ft, prog, i', pos, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   389
  from stimes_sgD [OF this[unfolded st_def], unfolded trset_of.simps]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   390
  have "tpc_set i \<subseteq> tprog_set prog \<union> tpc_set i' \<union> tpos_set pos \<union>  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   391
            tmem_set mem \<union> tfaults_set ft" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   392
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   393
    by (unfold tpn_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   394
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   395
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   396
lemma psD: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   397
  "(ps p \<and>* r) (trset_of (ft, prog, i', pos, mem)) \<Longrightarrow> pos = p"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   398
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   399
  assume "(ps p ** r) (trset_of (ft, prog, i', pos, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   400
  from stimes_sgD [OF this[unfolded ps_def], unfolded trset_of.simps]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   401
  have "tpos_set p \<subseteq> tprog_set prog \<union> tpc_set i' \<union> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   402
                       tpos_set pos \<union> tmem_set mem \<union> tfaults_set ft"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   403
    by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   404
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   405
    by (unfold tpn_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   406
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   407
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   408
lemma codeD: "(st i \<and>* sg {TC i inst} \<and>* r) (trset_of (ft, prog, i', pos, mem))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   409
       \<Longrightarrow> prog $ i = Some inst"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   410
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   411
  assume "(st i \<and>* sg {TC i inst} \<and>* r) (trset_of (ft, prog, i', pos, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   412
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   413
    apply(unfold sep_conj_def set_ins_def sg_def trset_of.simps tpn_set_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   414
    by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   415
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   416
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   417
lemma memD: "((tm a v) \<and>* r) (trset_of (ft, prog, i, pos, mem))  \<Longrightarrow> mem $ a = Some v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   418
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   419
  assume "((tm a v) \<and>* r) (trset_of (ft, prog, i, pos, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   420
  from stimes_sgD[OF this[unfolded trset_of.simps tpn_set_def tm_def]]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   421
  have "{TM a v} \<subseteq> {TC i inst |i inst. prog $ i = Some inst} \<union> {TAt i} \<union> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   422
    {TPos pos} \<union> {TM i n |i n. mem $ i = Some n} \<union> {TFaults ft}" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   423
  thus ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   424
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   425
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   426
lemma t_hoare_seq: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   427
  assumes a1: "\<And> i j. \<lbrace>st i \<and>* p\<rbrace> i:[c1]:j \<lbrace>st j \<and>* q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   428
  and     a2: "\<And> j k. \<lbrace>st j ** q\<rbrace> j:[c2]:k \<lbrace>st k ** r\<rbrace>" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   429
  shows "\<lbrace>st i \<and>* p\<rbrace> i:[(c1 ; c2)]:k \<lbrace>st k \<and>* r\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   430
proof(subst tassemble_to.simps, rule tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   431
  fix j'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   432
  show "\<lbrace>st i \<and>* p\<rbrace>  i:[ c1 ]:j' \<and>* j':[ c2 ]:k \<lbrace>st k \<and>* r\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   433
  proof(rule tm.composition)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   434
    from a1 show "\<lbrace>st i \<and>* p\<rbrace>  i:[ c1 ]:j' \<lbrace>st j' \<and>* q\<rbrace>" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   435
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   436
    from a2 show "\<lbrace>st j' \<and>* q\<rbrace>  j':[ c2 ]:k \<lbrace>st k \<and>* r\<rbrace>" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   437
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   438
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   439
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   440
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   441
lemma t_hoare_seq1:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
  assumes a1: "\<And>j'. \<lbrace>st i \<and>* p\<rbrace> i:[c1]:j' \<lbrace>st j' \<and>* q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443
  assumes a2: "\<And>j'. \<lbrace>st j' \<and>* q\<rbrace> j':[c2]:k \<lbrace>st k' \<and>* r\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   444
  shows "\<lbrace>st i \<and>* p\<rbrace> i:[(c1 ; c2)]:k \<lbrace>st k' \<and>* r\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   445
proof(subst tassemble_to.simps, rule tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   446
  fix j'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   447
  show "\<lbrace>st i \<and>* p\<rbrace>  i:[ c1 ]:j' \<and>* j':[ c2 ]:k \<lbrace>st k' \<and>* r\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   448
  proof(rule tm.composition)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   449
    from a1 show "\<lbrace>st i \<and>* p\<rbrace>  i:[ c1 ]:j' \<lbrace>st j' \<and>* q\<rbrace>" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   450
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   451
    from a2 show " \<lbrace>st j' \<and>* q\<rbrace>  j':[ c2 ]:k \<lbrace>st k' \<and>* r\<rbrace>" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   452
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   453
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   454
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   455
lemma t_hoare_seq2:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   456
  assumes h: "\<And>j. \<lbrace>st i ** p\<rbrace> i:[c1]:j \<lbrace>st k' \<and>* r\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   457
  shows "\<lbrace>st i ** p\<rbrace> i:[(c1 ; c2)]:j \<lbrace>st k' ** r\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   458
  apply (unfold tassemble_to.simps, rule tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   459
  by (rule tm.code_extension, rule h)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   460
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   461
lemma t_hoare_local: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   462
  assumes h: "(\<And>l. \<lbrace>st i \<and>* p\<rbrace>  i :[ c l ]: j \<lbrace>st k \<and>* q\<rbrace>)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   463
  shows "\<lbrace>st i ** p\<rbrace> i:[TLocal c]:j \<lbrace>st k ** q\<rbrace>" using h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   464
  by (unfold tassemble_to.simps, intro tm.code_exI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   465
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   466
lemma t_hoare_label: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   467
  assumes "\<And>l. l = i \<Longrightarrow> \<lbrace>st l \<and>* p\<rbrace>  l:[ c l ]:j \<lbrace>st k \<and>* q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   468
  shows "\<lbrace>st i \<and>* p\<rbrace> i:[(TLabel l; c l)]:j \<lbrace>st k \<and>* q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   469
using assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   470
by (unfold tassemble_to.simps, intro tm.code_exI tm.code_condI, clarify, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   471
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   472
primrec t_split_cmd :: "tpg \<Rightarrow> (tpg \<times> tpg option)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   473
  where "t_split_cmd (\<guillemotright>inst) = (\<guillemotright>inst, None)" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   474
        "t_split_cmd (TLabel l) = (TLabel l, None)" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   475
        "t_split_cmd (TSeq c1 c2) = (case (t_split_cmd c2) of
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   476
                                      (c21, Some c22) \<Rightarrow> (TSeq c1 c21, Some c22) |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   477
                                      (c21, None) \<Rightarrow> (c1, Some c21))" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   478
        "t_split_cmd (TLocal c) = (TLocal c, None)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   479
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   480
definition "t_last_cmd tpg = snd (t_split_cmd tpg)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   481
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   482
declare t_last_cmd_def [simp]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   483
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   484
definition "t_blast_cmd tpg = fst (t_split_cmd tpg)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   485
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   486
declare t_blast_cmd_def [simp]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   487
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   488
lemma "t_last_cmd (c1; c2; TLabel l) = Some (TLabel l)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   489
  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   490
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   491
lemma "t_blast_cmd (c1; c2; TLabel l) = (c1; c2)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   492
  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   493
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   494
lemma t_split_cmd_eq:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   495
  assumes "t_split_cmd c = (c1, Some c2)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   496
  shows "i:[c]:j = i:[(c1; c2)]:j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   497
  using assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   498
proof(induct c arbitrary: c1 c2 i j)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   499
  case (TSeq cx cy)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   500
  from `t_split_cmd (cx ; cy) = (c1, Some c2)`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   501
  show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   502
    apply (simp split:prod.splits option.splits)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   503
    apply (cases cy, auto split:prod.splits option.splits)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   504
    proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   505
      fix a
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   506
      assume h: "t_split_cmd cy = (a, Some c2)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   507
      show "i :[ (cx ; cy) ]: j = i :[ ((cx ; a) ; c2) ]: j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   508
        apply (simp only: tassemble_to.simps(2) TSeq(2)[OF h] sep_conj_exists)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   509
        apply (simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   510
        by (simp add:pred_ex_def, default, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   511
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   512
qed auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   513
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   514
lemma t_hoare_label_last_pre: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   515
  assumes h1: "t_split_cmd c = (c', Some (TLabel l))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   516
  and h2: "l = j \<Longrightarrow> \<lbrace>p\<rbrace> i:[c']:j \<lbrace>q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   517
  shows "\<lbrace>p\<rbrace> i:[c]:j \<lbrace>q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   518
by (unfold t_split_cmd_eq[OF h1], 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   519
    simp only:tassemble_to.simps sep_conj_cond, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   520
    intro tm.code_exI tm.code_condI, insert h2, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   521
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   522
lemma t_hoare_label_last: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   523
  assumes h1: "t_last_cmd c = Some (TLabel l)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   524
  and h2: "l = j \<Longrightarrow> \<lbrace>p\<rbrace> i:[t_blast_cmd c]:j \<lbrace>q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   525
  shows "\<lbrace>p\<rbrace> i:[c]:j \<lbrace>q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   526
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   527
    have "t_split_cmd c = (t_blast_cmd c, t_last_cmd c)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   528
      by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   529
  from t_hoare_label_last_pre[OF this[unfolded h1]] h2
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   530
  show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   531
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   532
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   533
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   534
subsection {* Basic assertions for TM *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   535
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   536
(* ones between tape position i and j *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   537
function ones :: "int \<Rightarrow> int \<Rightarrow> tassert" where
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   538
  "ones i j = (if j < i then <(i = j + 1)> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   539
               else (one i) \<and>* ones (i + 1) j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   540
by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   541
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   542
termination 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   543
  by (relation "measure(\<lambda> (i::int, j). nat (j - i + 1))") auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   544
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   545
lemma ones_base_simp: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   546
  "j < i \<Longrightarrow> ones i j = <(i = j + 1)>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   547
  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   548
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   549
lemma ones_step_simp: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   550
  "\<not> j < i \<Longrightarrow> ones i j =  ((one i) \<and>* ones (i + 1) j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   551
  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   552
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   553
lemmas ones_simps = ones_base_simp ones_step_simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   554
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   555
declare ones.simps [simp del] ones_simps [simp]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   556
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   557
lemma ones_induct [case_names Base Step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   558
  assumes h1: "\<And> i j. j < i \<Longrightarrow> P i j (<(i = j + (1::int))>)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   559
  and h2: "\<And> i j. \<lbrakk>\<not> j < i; P (i + 1) j (ones (i + 1) j)\<rbrakk> \<Longrightarrow> P i j ((one i) \<and>* ones (i + 1) j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   560
  shows "P i j (ones i j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   561
proof(induct i j rule:ones.induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   562
  fix i j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   563
  assume ih: "(\<not> j < i \<Longrightarrow> P (i + 1) j (ones (i + 1) j))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   564
  show "P i j (ones i j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   565
  proof(cases "j < i")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   566
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   567
    with h1 [OF True]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   568
    show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   569
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   570
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   571
    from h2 [OF False] and ih[OF False]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   572
    have "P i j (one i \<and>* ones (i + 1) j)" by blast
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   573
    with False show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   574
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   575
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   576
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   577
function ones' ::  "int \<Rightarrow> int \<Rightarrow> tassert" where
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   578
  "ones' i j = (if j < i then <(i = j + 1)> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   579
                else ones' i (j - 1) \<and>* one j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   580
by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   581
termination by (relation "measure(\<lambda> (i::int, j). nat (j - i + 1))") auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   582
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   583
lemma ones_rev: "\<not> j < i \<Longrightarrow> (ones i j) = ((ones i (j - 1)) ** one j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   584
proof(induct i j rule:ones_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   585
  case Base
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   586
  thus ?case by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   587
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   588
  case (Step i j)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   589
  show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   590
  proof(cases "j < i + 1")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   591
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   592
    with Step show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   593
      by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   594
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   595
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   596
    with Step show ?thesis 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   597
      by (auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   598
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   599
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   600
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   601
lemma ones_rev_induct [case_names Base Step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   602
  assumes h1: "\<And> i j. j < i \<Longrightarrow> P i j (<(i = j + (1::int))>)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   603
  and h2: "\<And> i j. \<lbrakk>\<not> j < i; P i (j - 1) (ones i (j - 1))\<rbrakk> \<Longrightarrow> P i j ((ones i (j - 1)) ** one j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   604
  shows "P i j (ones i j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   605
proof(induct i j rule:ones'.induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   606
  fix i j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   607
  assume ih: "(\<not> j < i \<Longrightarrow> P i (j - 1) (ones i (j - 1)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   608
  show "P i j (ones i j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   609
  proof(cases "j < i")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   610
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   611
    with h1 [OF True]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   612
    show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   613
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   614
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   615
    from h2 [OF False] and ih[OF False]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   616
    have "P i j (ones i (j - 1) \<and>* one j)" by blast
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   617
    with ones_rev and False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   618
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   619
      by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   620
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   621
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   622
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   623
function zeros :: "int \<Rightarrow> int \<Rightarrow> tassert" where
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   624
  "zeros i j = (if j < i then <(i = j + 1)> else
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   625
                (zero i) ** zeros (i + 1) j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   626
by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   627
termination by (relation "measure(\<lambda> (i::int, j). nat (j - i + 1))") auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   628
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   629
lemma zeros_base_simp: "j < i \<Longrightarrow> zeros i j = <(i = j + 1)>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   630
  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   631
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   632
lemma zeros_step_simp: "\<not> j < i \<Longrightarrow> zeros i j = ((zero i) ** zeros (i + 1) j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   633
  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   634
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   635
declare zeros.simps [simp del]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   636
lemmas zeros_simps [simp] = zeros_base_simp zeros_step_simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   637
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   638
lemma zeros_induct [case_names Base Step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   639
  assumes h1: "\<And> i j. j < i \<Longrightarrow> P i j (<(i = j + (1::int))>)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   640
  and h2: "\<And> i j. \<lbrakk>\<not> j < i; P (i + 1) j (zeros (i + 1) j)\<rbrakk> \<Longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   641
                                   P i j ((zero i) ** zeros (i + 1) j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   642
  shows "P i j (zeros i j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   643
proof(induct i j rule:zeros.induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   644
  fix i j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   645
  assume ih: "(\<not> j < i \<Longrightarrow> P (i + 1) j (zeros (i + 1) j))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   646
  show "P i j (zeros i j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   647
  proof(cases "j < i")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   648
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   649
    with h1 [OF True]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   650
    show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   651
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   652
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   653
    from h2 [OF False] and ih[OF False]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   654
    have "P i j (zero i \<and>* zeros (i + 1) j)" by blast
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   655
    with False show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   656
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   657
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   658
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   659
lemma zeros_rev: "\<not> j < i \<Longrightarrow> (zeros i j) = ((zeros i (j - 1)) \<and>* zero j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   660
proof(induct i j rule:zeros_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   661
  case (Base i j)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   662
  thus ?case by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   663
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   664
  case (Step i j)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   665
  show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   666
  proof(cases "j < i + 1")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   667
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   668
    with Step show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   669
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   670
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   671
    with Step show ?thesis by (auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   672
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   673
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   674
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   675
lemma zeros_rev_induct [case_names Base Step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   676
  assumes h1: "\<And> i j. j < i \<Longrightarrow> P i j (<(i = j + (1::int))>)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   677
  and h2: "\<And> i j. \<lbrakk>\<not> j < i; P i (j - 1) (zeros i (j - 1))\<rbrakk> \<Longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   678
                       P i j ((zeros i (j - 1)) ** zero j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   679
  shows "P i j (zeros i j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   680
proof(induct i j rule:ones'.induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   681
  fix i j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   682
  assume ih: "(\<not> j < i \<Longrightarrow> P i (j - 1) (zeros i (j - 1)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   683
  show "P i j (zeros i j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   684
  proof(cases "j < i")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   685
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   686
    with h1 [OF True]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   687
    show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   688
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   689
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   690
    from h2 [OF False] and ih[OF False]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   691
    have "P i j (zeros i (j - 1) \<and>* zero j)" by blast
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   692
    with zeros_rev and False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   693
    show ?thesis by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   694
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   695
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   696
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   697
definition "rep i j k = ((ones i (i + (int k))) \<and>* <(j = i + int k)>)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   698
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   699
fun reps :: "int \<Rightarrow> int \<Rightarrow> nat list\<Rightarrow> tassert"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   700
  where
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   701
  "reps i j [] = <(i = j + 1)>" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   702
  "reps i j [k] = (ones i (i + int k) ** <(j = i + int k)>)" |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   703
  "reps i j (k # ks) = (ones i (i + int k) ** zero (i + int k + 1) ** reps (i + int k + 2) j ks)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   704
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   705
lemma reps_simp3: "ks \<noteq> [] \<Longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   706
  reps i j (k # ks) = (ones i (i + int k) ** zero (i + int k + 1) ** reps (i + int k + 2) j ks)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   707
  by (cases ks, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   708
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   709
definition "reps_sep_len ks = (if length ks \<le> 1 then 0 else (length ks) - 1)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   710
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   711
definition "reps_ctnt_len ks = (\<Sum> k \<leftarrow> ks. (1 + k))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   712
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   713
definition "reps_len ks = (reps_sep_len ks) + (reps_ctnt_len ks)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   714
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   715
definition "splited xs ys zs = ((xs = ys @ zs) \<and> (ys \<noteq> []) \<and> (zs \<noteq> []))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   716
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   717
lemma list_split: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   718
  assumes h: "k # ks = ys @ zs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   719
      and h1: "ys \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   720
  shows "(ys = [k] \<and> zs = ks) \<or> (\<exists> ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> ks = ys' @ zs)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   721
proof(cases ys)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   722
  case Nil
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   723
  with h1 show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   724
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   725
  case (Cons y' ys')
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   726
  with h have "k#ks = y'#(ys' @ zs)" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   727
  hence hh: "y' = k" "ks = ys' @ zs" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   728
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   729
  proof(cases "ys' = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   730
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   731
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   732
    proof(rule disjI2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   733
      show " \<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> ks = ys' @ zs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   734
      proof(rule exI[where x = ys'])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   735
        from False hh Cons show "ys' \<noteq> [] \<and> ys = k # ys' \<and> ks = ys' @ zs" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   736
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   737
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   738
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   739
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   740
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   741
    proof(rule disjI1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   742
      from hh True Cons
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   743
      show "ys = [k] \<and> zs = ks" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   744
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   745
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   746
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   747
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   748
lemma splited_cons[elim_format]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   749
  assumes h: "splited (k # ks) ys zs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   750
  shows "(ys = [k] \<and> zs = ks) \<or> (\<exists> ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> splited ks ys' zs)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   751
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   752
  from h have "k # ks = ys @ zs" "ys \<noteq> [] " unfolding splited_def by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   753
  from list_split[OF this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   754
  have "ys = [k] \<and> zs = ks \<or> (\<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> ks = ys' @ zs)" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   755
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   756
  proof
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   757
    assume " ys = [k] \<and> zs = ks" thus ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   758
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   759
    assume "\<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> ks = ys' @ zs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   760
    then obtain ys' where hh: "ys' \<noteq> []" "ys = k # ys'" "ks = ys' @ zs" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   761
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   762
    proof(rule disjI2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   763
      show "\<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> splited ks ys' zs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   764
      proof(rule exI[where x = ys'])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   765
        from h have "zs \<noteq> []" by (unfold splited_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   766
        with hh(1) hh(3)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   767
        have "splited ks ys' zs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   768
          by (unfold splited_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   769
        with hh(1) hh(2) show "ys' \<noteq> [] \<and> ys = k # ys' \<and> splited ks ys' zs" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   770
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   771
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   772
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   773
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   774
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   775
lemma splited_cons_elim:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   776
  assumes h: "splited (k # ks) ys zs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   777
  and h1: "\<lbrakk>ys = [k]; zs = ks\<rbrakk> \<Longrightarrow> P"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   778
  and h2: "\<And> ys'. \<lbrakk>ys' \<noteq> []; ys = k#ys'; splited ks ys' zs\<rbrakk> \<Longrightarrow> P"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   779
  shows P
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   780
proof(rule splited_cons[OF h])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   781
  assume "ys = [k] \<and> zs = ks \<or> (\<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> splited ks ys' zs)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   782
  thus P
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   783
  proof
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   784
    assume hh: "ys = [k] \<and> zs = ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   785
    show P
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   786
    proof(rule h1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   787
      from hh show "ys = [k]" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   788
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   789
      from hh show "zs = ks" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   790
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   791
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   792
    assume "\<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> splited ks ys' zs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   793
    then obtain ys' where hh: "ys' \<noteq> []" "ys = k # ys'"  "splited ks ys' zs" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   794
    from h2[OF this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   795
    show P .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   796
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   797
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   798
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   799
lemma list_nth_expand:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   800
  "i < length xs \<Longrightarrow> xs = take i xs @ [xs!i] @ drop (Suc i) xs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   801
  by (metis Cons_eq_appendI append_take_drop_id drop_Suc_conv_tl eq_Nil_appendI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   802
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   803
lemma reps_len_nil: "reps_len [] = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   804
   by (auto simp:reps_len_def reps_sep_len_def reps_ctnt_len_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   805
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   806
lemma reps_len_sg: "reps_len [k] = 1 + k"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   807
  by (auto simp:reps_len_def reps_sep_len_def reps_ctnt_len_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   808
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   809
lemma reps_len_cons: "ks \<noteq> [] \<Longrightarrow> (reps_len (k # ks)) = 2 + k + reps_len ks "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   810
proof(induct ks arbitrary:k)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   811
  case (Cons n ns)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   812
  thus ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   813
    by(cases "ns = []", 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   814
      auto simp:reps_len_def reps_sep_len_def reps_ctnt_len_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   815
qed auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   816
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   817
lemma reps_len_correct:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   818
  assumes h: "(reps i j ks) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   819
  shows "j = i + int (reps_len ks) - 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   820
  using h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   821
proof(induct ks arbitrary:i j s)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   822
  case Nil
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   823
  thus ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   824
    by (auto simp:reps_len_nil pasrt_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   825
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   826
  case (Cons n ns)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   827
  thus ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   828
  proof(cases "ns = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   829
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   830
    from Cons(2)[unfolded reps_simp3[OF False]]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   831
    obtain s' where "(reps (i + int n + 2) j ns) s'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   832
      by (auto elim!:sep_conjE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   833
    from Cons.hyps[OF this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   834
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   835
      by (unfold reps_len_cons[OF False], simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   836
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   837
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   838
    with Cons(2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   839
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   840
      by (auto simp:reps_len_sg pasrt_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   841
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   842
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   843
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   844
lemma reps_len_expand: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   845
  shows "(EXS j. (reps i j ks)) = (reps i (i + int (reps_len ks) - 1) ks)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   846
proof(default+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   847
  fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   848
  assume "(EXS j. reps i j ks) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   849
  with reps_len_correct show "reps i (i + int (reps_len ks) - 1) ks s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   850
    by (auto simp:pred_ex_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   851
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   852
  fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   853
  assume "reps i (i + int (reps_len ks) - 1) ks s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   854
  thus "(EXS j. reps i j ks) s"  by (auto simp:pred_ex_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   855
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   856
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   857
lemma reps_len_expand1: "(EXS j. (reps i j ks \<and>* r)) = (reps i (i + int (reps_len ks) - 1) ks \<and>* r)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   858
by (metis reps_len_def reps_len_expand sep.mult_commute sep_conj_exists1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   859
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   860
lemma reps_splited:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   861
  assumes h: "splited xs ys zs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   862
  shows "reps i j xs = (reps i (i + int (reps_len ys) - 1) ys \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   863
                        zero (i + int (reps_len ys)) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   864
                        reps (i + int (reps_len ys) + 1) j zs)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   865
  using h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   866
proof(induct xs arbitrary: i j ys zs)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   867
  case Nil
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   868
  thus ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   869
    by (unfold splited_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   870
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   871
  case (Cons k ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   872
  from Cons(2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   873
  show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   874
  proof(rule splited_cons_elim)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   875
    assume h: "ys = [k]" "zs = ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   876
    with Cons(2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   877
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   878
    proof(cases "ks = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   879
      case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   880
      with Cons(2) have False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   881
        by (simp add:splited_def, cases ys, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   882
      thus ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   883
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   884
      case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   885
      have ss: "i + int k + 1 = i + (1 + int k)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   886
           "i + int k + 2 = 2 + (i + int k)" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   887
      show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   888
        by (unfold h(1), unfold reps_simp3[OF False] reps.simps(2) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   889
            reps_len_sg, auto simp:sep_conj_ac,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   890
            unfold ss h(2), simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   891
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   892
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   893
    fix ys'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   894
    assume h: "ys' \<noteq> []" "ys = k # ys'" "splited ks ys' zs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   895
    hence nnks: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   896
      by (unfold splited_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   897
    have ss: "i + int k + 2 + int (reps_len ys') = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   898
              i + (2 + (int k + int (reps_len ys')))" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   899
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   900
      by (simp add: reps_simp3[OF nnks] reps_simp3[OF h(1)] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   901
                    Cons(1)[OF h(3)] h(2) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   902
                    reps_len_cons[OF h(1)]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   903
                    sep_conj_ac, subst ss, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   904
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   905
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   906
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   907
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   908
subsection {* A theory of list extension *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   909
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   910
definition "list_ext n xs = xs @ replicate ((Suc n) - length xs) 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   911
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   912
lemma list_ext_len_eq: "length (list_ext a xs) = length xs + (Suc a - length xs)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   913
  by (metis length_append length_replicate list_ext_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   914
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   915
lemma list_ext_len: "a < length (list_ext a xs)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   916
  by (unfold list_ext_len_eq, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   917
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   918
lemma list_ext_lt: "a < length xs \<Longrightarrow> list_ext a xs = xs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   919
  by (smt append_Nil2 list_ext_def replicate_0)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   920
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   921
lemma list_ext_lt_get: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   922
  assumes h: "a' < length xs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   923
  shows "(list_ext a xs)!a' = xs!a'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   924
proof(cases "a < length xs")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   925
  case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   926
  with h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   927
  show ?thesis by (auto simp:list_ext_lt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   928
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   929
  case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   930
  with h show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   931
    apply (unfold list_ext_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   932
    by (metis nth_append)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   933
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   934
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   935
lemma list_ext_get_upd: "((list_ext a xs)[a:=v])!a = v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   936
  by (metis list_ext_len nth_list_update_eq)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   937
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   938
lemma nth_app: "length xs \<le> a \<Longrightarrow> (xs @ ys)!a = ys!(a - length xs)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   939
  by (metis not_less nth_append)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   940
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   941
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   942
lemma list_ext_tail:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   943
  assumes h1: "length xs \<le> a"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   944
  and h2: "length xs \<le> a'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   945
  and h3: "a' \<le> a"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   946
  shows "(list_ext a xs)!a' = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   947
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   948
  from h1 h2
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   949
  have "a' - length xs < length (replicate (Suc a - length xs) 0)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   950
    by (metis diff_less_mono h3 le_imp_less_Suc length_replicate)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   951
  moreover from h1 have "replicate (Suc a - length xs) 0 \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   952
    by (smt empty_replicate)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   953
  ultimately have "(replicate (Suc a - length xs) 0)!(a' - length xs) = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   954
    by (metis length_replicate nth_replicate)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   955
  moreover have "(xs @ (replicate (Suc a - length xs) 0))!a' = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   956
            (replicate (Suc a - length xs) 0)!(a' - length xs)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   957
    by (rule nth_app[OF h2])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   958
  ultimately show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   959
    by (auto simp:list_ext_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   960
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   961
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   962
lemmas list_ext_simps = list_ext_lt_get list_ext_lt list_ext_len list_ext_len_eq
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   963
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   964
lemma listsum_upd_suc:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   965
  "a < length ks \<Longrightarrow> listsum (map Suc (ks[a := Suc (ks ! a)]))= (Suc (listsum (map Suc ks)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   966
by (smt elem_le_listsum_nat 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   967
     length_list_update list_ext_get_upd 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   968
     list_update_overwrite listsum_update_nat map_update 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   969
     nth_equalityI nth_list_update nth_map)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   970
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   971
lemma reps_len_suc:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   972
  assumes "a < length ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   973
  shows "reps_len (ks[a:=Suc(ks!a)]) = 1 + reps_len ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   974
proof(cases "length ks \<le> 1")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   975
  case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   976
  with `a < length ks` 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   977
  obtain k where "ks = [k]" "a = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   978
    by (smt length_0_conv length_Suc_conv)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   979
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   980
      apply (unfold `ks = [k]` `a = 0`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   981
      by (unfold reps_len_def reps_sep_len_def reps_ctnt_len_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   982
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   983
  case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   984
  have "Suc = (op + (Suc 0))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   985
    by (default, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   986
  with listsum_upd_suc[OF `a < length ks`] False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   987
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   988
     by (unfold reps_len_def reps_sep_len_def reps_ctnt_len_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   989
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   990
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   991
lemma ks_suc_len:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   992
  assumes h1: "(reps i j ks) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   993
  and h2: "ks!a = v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   994
  and h3: "a < length ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   995
  and h4: "(reps i j' (ks[a:=Suc v])) s'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   996
  shows "j' = j + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   997
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   998
  from reps_len_correct[OF h1, unfolded list_ext_len_eq]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   999
       reps_len_correct[OF h4, unfolded list_ext_len_eq] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1000
       reps_len_suc[OF `a < length ks`] h2 h3
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1001
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1002
    by (unfold list_ext_lt[OF `a < length ks`], auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1003
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1004
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1005
lemma ks_ext_len:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1006
  assumes h1: "(reps i j ks) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1007
  and h3: "length ks \<le> a"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1008
  and h4: "reps i j' (list_ext a ks) s'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1009
  shows "j' = j + int (reps_len (list_ext a ks)) - int (reps_len ks)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1010
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1011
  from reps_len_correct[OF h1, unfolded  list_ext_len_eq]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1012
    and reps_len_correct[OF h4, unfolded list_ext_len_eq]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1013
  h3
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1014
  show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1015
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1016
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1017
definition 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1018
  "reps' i j ks = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1019
     (if ks = [] then (<(i = j + 1)>)  else (reps i (j - 1) ks \<and>* zero j))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1020
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1021
lemma reps'_expand: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1022
  assumes h: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1023
  shows "(EXS j. reps' i j ks) = reps' i (i + int (reps_len ks)) ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1024
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1025
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1026
  proof(default+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1027
    fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1028
    assume "(EXS j. reps' i j ks) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1029
    with h have "(EXS j. reps i (j - 1) ks \<and>* zero j) s" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1030
      by (simp add:reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1031
    hence "(reps i (i + int (reps_len ks) - 1) ks \<and>* zero (i + int (reps_len ks))) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1032
    proof(elim EXS_elim)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1033
      fix j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1034
      assume "(reps i (j - 1) ks \<and>* zero j) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1035
      then obtain s1 s2 where "s = s1 + s2" "s1 ## s2" "reps i (j - 1) ks s1" "zero j s2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1036
        by (auto elim!:sep_conjE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1037
      from reps_len_correct[OF this(3)]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1038
      have "j = i + int (reps_len ks)" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1039
      with `reps i (j - 1) ks s1` have "reps i (i + int (reps_len ks) - 1) ks s1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1040
        by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1041
      moreover from `j = i + int (reps_len ks)` and `zero j s2`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1042
      have "zero (i + int (reps_len ks)) s2" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1043
      ultimately show "(reps i (i + int (reps_len ks) - 1) ks \<and>* zero (i + int (reps_len ks))) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1044
        using `s = s1 + s2` `s1 ## s2` by (auto intro!:sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1045
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1046
    thus "reps' i (i + int (reps_len ks)) ks s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1047
      by (simp add:h reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1048
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1049
    fix s 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1050
    assume "reps' i (i + int (reps_len ks)) ks s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1051
    thus "(EXS j. reps' i j ks) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1052
      by (auto intro!:EXS_intro)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1053
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1054
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1055
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1056
lemma reps'_len_correct: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1057
  assumes h: "(reps' i j ks) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1058
  and h1: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1059
  shows "(j = i + int (reps_len ks))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1060
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1061
  from h1 have "reps' i j ks s = (reps i (j - 1) ks \<and>* zero j) s" by (simp add:reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1062
  from h[unfolded this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1063
  obtain s' where "reps i (j - 1) ks s'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1064
    by (auto elim!:sep_conjE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1065
  from reps_len_correct[OF this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1066
  show ?thesis by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1067
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1068
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1069
lemma reps'_append:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1070
  "reps' i j (ks1 @ ks2) = (EXS m. (reps' i (m - 1) ks1 \<and>* reps' m j ks2))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1071
proof(cases "ks1 = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1072
  case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1073
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1074
    by (auto simp:reps'_def pred_ex_def pasrt_def set_ins_def sep_conj_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1075
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1076
  case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1077
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1078
  proof(cases "ks2 = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1079
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1080
    from `ks1 \<noteq> []` and `ks2 \<noteq> []` 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1081
    have "splited (ks1 @ ks2) ks1 ks2" by (auto simp:splited_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1082
    from reps_splited[OF this, of i "j - 1"]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1083
    have eq_1: "reps i (j - 1) (ks1 @ ks2) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1084
           (reps i (i + int (reps_len ks1) - 1) ks1 \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1085
           zero (i + int (reps_len ks1)) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1086
           reps (i + int (reps_len ks1) + 1) (j - 1) ks2)" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1087
    from `ks1 \<noteq> []`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1088
    have eq_2: "reps' i j (ks1 @ ks2) = (reps i (j - 1) (ks1 @ ks2) \<and>* zero j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1089
      by (unfold reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1090
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1091
    proof(default+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1092
      fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1093
      assume "reps' i j (ks1 @ ks2) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1094
      show "(EXS m. reps' i (m - 1) ks1 \<and>* reps' m j ks2) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1095
      proof(rule EXS_intro[where x = "i + int(reps_len ks1) + 1"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1096
        from `reps' i j (ks1 @ ks2) s`[unfolded eq_2 eq_1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1097
        and `ks1 \<noteq> []` `ks2 \<noteq> []`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1098
        show "(reps' i (i + int (reps_len ks1) + 1 - 1) ks1 \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1099
                         reps' (i + int (reps_len ks1) + 1) j ks2) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1100
          by (unfold reps'_def, simp, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1101
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1102
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1103
      fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1104
      assume "(EXS m. reps' i (m - 1) ks1 \<and>* reps' m j ks2) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1105
      thus "reps' i j (ks1 @ ks2) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1106
      proof(elim EXS_elim)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1107
        fix m
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1108
        assume "(reps' i (m - 1) ks1 \<and>* reps' m j ks2) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1109
        then obtain s1 s2 where h: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1110
          "s = s1 + s2" "s1 ## s2" "reps' i (m - 1) ks1 s1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1111
          " reps' m j ks2 s2" by (auto elim!:sep_conjE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1112
        from reps'_len_correct[OF this(3) `ks1 \<noteq> []`]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1113
        have eq_m: "m = i + int (reps_len ks1) + 1" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1114
        have "((reps i (i + int (reps_len ks1) - 1) ks1 \<and>* zero (i + int (reps_len ks1))) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1115
               ((reps (i + int (reps_len ks1) + 1) (j - 1) ks2) \<and>* zero j)) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1116
          (is "(?P \<and>* ?Q) s") 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1117
        proof(rule sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1118
          from h(3) eq_m `ks1 \<noteq> []` show "?P s1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1119
            by (simp add:reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1120
        next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1121
          from h(4) eq_m `ks2 \<noteq> []` show "?Q s2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1122
            by (simp add:reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1123
        next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1124
          from h(2) show "s1 ## s2" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1125
        next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1126
          from h(1) show "s = s1 + s2" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1127
        qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1128
        thus "reps' i j (ks1 @ ks2) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1129
          by (unfold eq_2 eq_1, auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1130
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1131
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1132
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1133
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1134
    have "-1 + j = j - 1" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1135
    with `ks1 \<noteq> []` True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1136
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1137
      apply (auto simp:reps'_def pred_ex_def pasrt_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1138
      apply (unfold `-1 + j = j - 1`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1139
      by (fold sep_empty_def, simp only:sep_conj_empty)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1140
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1141
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1142
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1143
lemma reps'_sg: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1144
  "reps' i j [k] = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1145
       (<(j = i + int k + 1)> \<and>* ones i (i + int k) \<and>* zero j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1146
  apply (unfold reps'_def, default, auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1147
  by (sep_cancel+, simp add:pasrt_def)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1148
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1149
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1150
section {* Basic macros for TM *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1151
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1152
definition "write_zero = (TL exit. \<guillemotright>((W0, exit), (W0, exit)); TLabel exit)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1153
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1154
lemma st_upd: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1155
  assumes pre: "(st i' ** r) (trset_of (f, x, i, y, z))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1156
  shows "(st i'' ** r) (trset_of (f, x,  i'', y, z))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1157
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1158
  from stimes_sgD[OF pre[unfolded st_def], unfolded trset_of.simps, unfolded stD[OF pre]]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1159
  have "r (tprog_set x \<union> tpc_set i' \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f - tpc_set i')"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1160
    by blast
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1161
  moreover have 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1162
    "(tprog_set x \<union> tpc_set i' \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f - tpc_set i') =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1163
     (tprog_set x \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1164
    by (unfold tpn_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1165
  ultimately have r_rest: "r (tprog_set x \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1166
    by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1167
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1168
  proof(rule sep_conjI[where Q = r, OF _ r_rest])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1169
    show "st i'' (tpc_set i'')" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1170
      by (unfold st_def sg_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1171
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1172
    show "tpc_set i'' ## tprog_set x \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1173
      by (unfold tpn_set_def sep_disj_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1174
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1175
    show "trset_of (f, x, i'', y, z) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1176
             tpc_set i'' + (tprog_set x \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1177
      by (unfold trset_of.simps plus_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1178
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1179
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1180
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1181
lemma pos_upd: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1182
  assumes pre: "(ps i ** r) (trset_of (f, x, y, i', z))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1183
  shows "(ps j ** r) (trset_of (f, x, y, j, z))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1184
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1185
  from stimes_sgD[OF pre[unfolded ps_def], unfolded trset_of.simps, unfolded psD[OF pre]]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1186
  have "r (tprog_set x \<union> tpc_set y \<union> tpos_set i \<union> tmem_set z \<union> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1187
              tfaults_set f - tpos_set i)" (is "r ?xs")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1188
    by blast
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1189
  moreover have 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1190
    "?xs = tprog_set x \<union> tpc_set y  \<union> tmem_set z \<union> tfaults_set f"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1191
    by (unfold tpn_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1192
  ultimately have r_rest: "r \<dots>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1193
    by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1194
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1195
  proof(rule sep_conjI[where Q = r, OF _ r_rest])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1196
    show "ps j (tpos_set j)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1197
      by (unfold ps_def sg_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1198
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1199
    show "tpos_set j ## tprog_set x \<union> tpc_set y \<union> tmem_set z \<union> tfaults_set f"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1200
      by (unfold tpn_set_def sep_disj_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1201
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1202
    show "trset_of (f, x, y, j, z) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1203
             tpos_set j + (tprog_set x \<union> tpc_set y \<union> tmem_set z \<union> tfaults_set f)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1204
      by (unfold trset_of.simps plus_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1205
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1206
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1207
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1208
lemma TM_in_simp: "({TM a v} \<subseteq> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1209
                      tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tmem_set mem \<union> tfaults_set f) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1210
                             ({TM a v} \<subseteq> tmem_set mem)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1211
  by (unfold tpn_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1212
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1213
lemma tmem_set_upd: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1214
  "{TM a v} \<subseteq> tmem_set mem \<Longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1215
        tmem_set (mem(a f\<mapsto> v')) = ((tmem_set mem) - {TM a v}) \<union> {TM a v'}"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1216
apply(unfold tpn_set_def) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1217
apply(auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1218
apply (metis the.simps the_lookup_fmap_upd the_lookup_fmap_upd_other)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1219
apply (metis the_lookup_fmap_upd_other)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1220
by (metis option.inject the_lookup_fmap_upd_other)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1221
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1222
lemma tmem_set_disj: "{TM a v} \<subseteq> tmem_set mem \<Longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1223
                            {TM a v'} \<inter>  (tmem_set mem - {TM a v}) = {}"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1224
  by (unfold tpn_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1225
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1226
lemma smem_upd: "((tm a v) ** r) (trset_of (f, x, y, z, mem))  \<Longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1227
                    ((tm a v') ** r) (trset_of (f, x, y, z, mem(a f\<mapsto> v')))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1228
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1229
  have eq_s: "(tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tmem_set mem \<union> tfaults_set f - {TM a v}) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1230
    (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1231
    by (unfold tpn_set_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1232
  assume g: "(tm a v \<and>* r) (trset_of (f, x, y, z, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1233
  from this[unfolded trset_of.simps tm_def]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1234
  have h: "(sg {TM a v} \<and>* r) (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tmem_set mem \<union> tfaults_set f)" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1235
  hence h0: "r (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1236
    by(sep_drule stimes_sgD, clarify, unfold eq_s, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1237
  from h TM_in_simp have "{TM a v} \<subseteq> tmem_set mem"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1238
    by(sep_drule stimes_sgD, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1239
  from tmem_set_upd [OF this] tmem_set_disj[OF this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1240
  have h2: "tmem_set (mem(a f\<mapsto> v')) = {TM a v'} \<union> (tmem_set mem - {TM a v})" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1241
           "{TM a v'} \<inter> (tmem_set mem - {TM a v}) = {}" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1242
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1243
  proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1244
    have "(tm a v' ** r) (tmem_set (mem(a f\<mapsto> v')) \<union> tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tfaults_set f)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1245
    proof(rule sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1246
      show "tm a v' ({TM a v'})" by (unfold tm_def sg_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1247
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1248
      from h0 show "r (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1249
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1250
      show "{TM a v'} ## tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1251
      proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1252
        from g have " mem $ a = Some v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1253
          by(sep_frule memD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1254
        thus "?thesis"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1255
          by(unfold tpn_set_def set_ins_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1256
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1257
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1258
      show "tmem_set (mem(a f\<mapsto> v')) \<union> tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tfaults_set f =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1259
    {TM a v'} + (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1260
        by (unfold h2(1) set_ins_def eq_s, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1261
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1262
    thus ?thesis 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1263
      apply (unfold trset_of.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1264
      by (metis sup_commute sup_left_commute)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1265
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1266
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1267
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1268
lemma hoare_write_zero: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1269
  "\<lbrace>st i ** ps p ** tm p v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1270
     i:[write_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1271
   \<lbrace>st j ** ps p ** tm p Bk\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1272
proof(unfold write_zero_def, intro t_hoare_local, rule t_hoare_label_last, simp, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1273
  show "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace>  i :[ \<guillemotright> ((W0, j), W0, j) ]: j \<lbrace>st j \<and>* ps p \<and>* tm p Bk\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1274
  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1275
        intro tm.code_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1276
    assume eq_j: "j = Suc i"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1277
    show "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace>  sg {TC i ((W0, Suc i), W0, Suc i)} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1278
          \<lbrace>st (Suc i) \<and>* ps p \<and>* tm p Bk\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1279
    proof(fold eq_j, unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1280
      fix ft prog cs i' mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1281
      assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* sg {TC i ((W0, j), W0, j)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1282
              (trset_of (ft, prog, cs, i', mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1283
      from h have "prog $ i = Some ((W0, j), W0, j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1284
        apply(rule_tac r = "r \<and>* ps p \<and>* tm p v" in codeD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1285
        by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1286
      from h and this have stp:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1287
        "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, j, i', mem(i' f\<mapsto> Bk))" (is "?x = ?y")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1288
        apply(sep_frule psD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1289
        apply(sep_frule stD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1290
        apply(sep_frule memD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1291
        by(cases v, unfold tm.run_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1292
      from h have "i' = p"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1293
        by(sep_drule psD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1294
      with h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1295
      have "(r \<and>* ps p \<and>* st j \<and>* tm p Bk \<and>* sg {TC i ((W0, j), W0, j)}) (trset_of ?x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1296
        apply(unfold stp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1297
        apply(sep_drule pos_upd, sep_drule st_upd, sep_drule smem_upd)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1298
        apply(auto simp: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1299
        done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1300
      thus "\<exists>k. (r \<and>* ps p \<and>* st j \<and>* tm p Bk \<and>* sg {TC i ((W0, j), W0, j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1301
             (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1302
        apply (rule_tac x = 0 in exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1303
        by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1304
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1305
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1306
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1307
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1308
lemma hoare_write_zero_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1309
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1310
  shows  "\<lbrace>st i ** ps p ** tm q v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1311
            i:[write_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1312
          \<lbrace>st j ** ps p ** tm q Bk\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1313
  by (unfold assms, rule hoare_write_zero)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1314
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1315
definition "write_one = (TL exit. \<guillemotright>((W1, exit), (W1, exit)); TLabel exit)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1316
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1317
lemma hoare_write_one: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1318
  "\<lbrace>st i ** ps p ** tm p v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1319
     i:[write_one]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1320
   \<lbrace>st j ** ps p ** tm p Oc\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1321
proof(unfold write_one_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1322
  fix l
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1323
  show "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace>  i :[ \<guillemotright> ((W1, j), W1, j) ]: j \<lbrace>st j \<and>* ps p \<and>* tm p Oc\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1324
  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1325
        rule_tac tm.code_condI, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1326
    let ?j = "Suc i"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1327
    show "\<lbrace>ps p \<and>* st i \<and>* tm p v\<rbrace>  sg {TC i ((W1, ?j), W1, ?j)} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1328
          \<lbrace>ps p \<and>* st ?j \<and>* tm p Oc\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1329
    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1330
      fix ft prog cs i' mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1331
      assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* sg {TC i ((W1, ?j), W1, ?j)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1332
              (trset_of (ft, prog, cs, i', mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1333
      from h have "prog $ i = Some ((W1, ?j), W1, ?j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1334
        apply(rule_tac r = "r \<and>* ps p \<and>* tm p v" in codeD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1335
        by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1336
      from h and this have stp:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1337
        "tm.run 1 (ft, prog, cs, i', mem) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1338
                     (ft, prog, ?j, i', mem(i' f\<mapsto> Oc))" (is "?x = ?y")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1339
        apply(sep_frule psD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1340
        apply(sep_frule stD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1341
        apply(sep_frule memD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1342
        by(cases v, unfold tm.run_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1343
      from h have "i' = p"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1344
        by(sep_drule psD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1345
      with h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1346
      have "(r \<and>* ps p \<and>* st ?j \<and>* tm p Oc \<and>* sg {TC i ((W1, ?j), W1, ?j)}) (trset_of ?x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1347
        apply(unfold stp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1348
        apply(sep_drule pos_upd, sep_drule st_upd, sep_drule smem_upd)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1349
        apply(auto simp: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1350
        done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1351
      thus "\<exists>k. (r \<and>* ps p \<and>* st ?j \<and>* tm p Oc \<and>* sg {TC i ((W1, ?j), W1, ?j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1352
             (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1353
        apply (rule_tac x = 0 in exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1354
        by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1355
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1356
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1357
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1358
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1359
lemma hoare_write_one_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1360
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1361
  shows  "\<lbrace>st i ** ps p ** tm q v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1362
              i:[write_one]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1363
          \<lbrace>st j ** ps p ** tm q Oc\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1364
  by (unfold assms, rule hoare_write_one)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1365
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1366
definition "move_left = (TL exit . \<guillemotright>((L, exit), (L, exit)); TLabel exit)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1367
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1368
lemma hoare_move_left: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1369
  "\<lbrace>st i ** ps p ** tm p v2\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1370
     i:[move_left]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1371
   \<lbrace>st j ** ps (p - 1) **  tm p v2\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1372
proof(unfold move_left_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1373
  fix l
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1374
  show "\<lbrace>st i \<and>* ps p \<and>* tm p v2\<rbrace>  i :[ \<guillemotright> ((L, l), L, l) ]: l
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1375
        \<lbrace>st l \<and>* ps (p - 1) \<and>* tm p v2\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1376
  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1377
      intro tm.code_condI, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1378
    let ?j = "Suc i"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1379
    show "\<lbrace>ps p \<and>* st i \<and>* tm p v2\<rbrace>  sg {TC i ((L, ?j), L, ?j)} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1380
          \<lbrace>st ?j \<and>* tm p v2 \<and>* ps (p - 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1381
    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1382
      fix ft prog cs i' mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1383
      assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v2 \<and>* sg {TC i ((L, ?j), L, ?j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1384
                       (trset_of (ft, prog, cs, i',  mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1385
      from h have "prog $ i = Some ((L, ?j), L, ?j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1386
        apply(rule_tac r = "r \<and>* ps p \<and>* tm p v2" in codeD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1387
        by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1388
      from h and this have stp:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1389
        "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, ?j, i' - 1, mem)" (is "?x = ?y")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1390
        apply(sep_frule psD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1391
        apply(sep_frule stD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1392
        apply(sep_frule memD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1393
        apply(unfold tm.run_def, case_tac v2, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1394
        done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1395
      from h have "i' = p"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1396
        by(sep_drule psD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1397
      with h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1398
      have "(r \<and>* st ?j \<and>* tm p v2 \<and>* ps (p - 1) \<and>* sg {TC i ((L, ?j), L, ?j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1399
               (trset_of ?x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1400
        apply(unfold stp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1401
        apply(sep_drule pos_upd, sep_drule st_upd, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1402
      proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1403
        assume "(st ?j \<and>* ps (p - 1) \<and>* r \<and>* tm p v2 \<and>* sg {TC i ((L, ?j), L, ?j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1404
                   (trset_of (ft, prog, ?j, p - 1, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1405
        thus "(r \<and>* st ?j \<and>* tm p v2 \<and>* ps (p - 1) \<and>* sg {TC i ((L, ?j), L, ?j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1406
                    (trset_of (ft, prog, ?j, p - 1, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1407
          by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1408
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1409
      thus "\<exists>k. (r \<and>* st ?j \<and>* tm p v2 \<and>* ps (p - 1) \<and>* sg {TC i ((L, ?j), L, ?j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1410
             (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1411
        apply (rule_tac x = 0 in exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1412
        by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1413
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1414
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1415
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1416
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1417
lemma hoare_move_left_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1418
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1419
  shows "\<lbrace>st i ** ps p ** tm q v2\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1420
            i:[move_left]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1421
         \<lbrace>st j ** ps (p - 1) **  tm q v2\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1422
  by (unfold assms, rule hoare_move_left)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1423
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1424
definition "move_right = (TL exit . \<guillemotright>((R, exit), (R, exit)); TLabel exit)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1425
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1426
lemma hoare_move_right: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1427
  "\<lbrace>st i ** ps p ** tm p v1 \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1428
     i:[move_right]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1429
   \<lbrace>st j ** ps (p + 1) ** tm p v1 \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1430
proof(unfold move_right_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1431
  fix l
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1432
  show "\<lbrace>st i \<and>* ps p \<and>* tm p v1\<rbrace>  i :[ \<guillemotright> ((R, l), R, l) ]: l
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1433
        \<lbrace>st l \<and>* ps (p + 1) \<and>* tm p v1\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1434
  proof(unfold tassemble_to.simps, simp only:sep_conj_cond, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1435
      intro tm.code_condI, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1436
    let ?j = "Suc i"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1437
    show "\<lbrace>ps p \<and>* st i \<and>* tm p v1\<rbrace>  sg {TC i ((R, ?j), R, ?j)} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1438
          \<lbrace>st ?j \<and>* tm p v1 \<and>* ps (p + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1439
    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1440
      fix ft prog cs i' mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1441
      assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v1 \<and>* sg {TC i ((R, ?j), R, ?j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1442
                       (trset_of (ft, prog, cs, i',  mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1443
      from h have "prog $ i = Some ((R, ?j), R, ?j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1444
        apply(rule_tac r = "r \<and>* ps p \<and>* tm p v1" in codeD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1445
        by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1446
      from h and this have stp:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1447
        "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, ?j, i'+ 1, mem)" (is "?x = ?y")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1448
        apply(sep_frule psD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1449
        apply(sep_frule stD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1450
        apply(sep_frule memD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1451
        apply(unfold tm.run_def, case_tac v1, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1452
        done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1453
      from h have "i' = p"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1454
        by(sep_drule psD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1455
      with h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1456
      have "(r \<and>* st ?j \<and>* tm p v1 \<and>* ps (p + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1457
                sg {TC i ((R, ?j), R, ?j)}) (trset_of ?x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1458
        apply(unfold stp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1459
        apply(sep_drule pos_upd, sep_drule st_upd, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1460
      proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1461
        assume "(st ?j \<and>* ps (p + 1) \<and>* r \<and>* tm p v1 \<and>* sg {TC i ((R, ?j), R, ?j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1462
                   (trset_of (ft, prog, ?j, p + 1, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1463
        thus "(r \<and>* st ?j \<and>* tm p v1 \<and>* ps (p + 1) \<and>* sg {TC i ((R, ?j), R, ?j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1464
                    (trset_of (ft, prog, ?j, p + 1, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1465
          by (simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1466
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1467
      thus "\<exists>k. (r \<and>* st ?j \<and>* tm p v1 \<and>* ps (p + 1) \<and>* sg {TC i ((R, ?j), R, ?j)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1468
             (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1469
        apply (rule_tac x = 0 in exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1470
        by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1471
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1472
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1473
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1474
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1475
lemma hoare_move_right_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1476
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1477
  shows "\<lbrace>st i ** ps p ** tm q v1 \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1478
           i:[move_right]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1479
         \<lbrace>st j ** ps (p + 1) ** tm q v1 \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1480
  by (unfold assms, rule hoare_move_right)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1481
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1482
definition "if_one e = (TL exit . \<guillemotright>((W0, exit), (W1, e)); TLabel exit)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1483
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1484
lemma hoare_if_one_true: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1485
  "\<lbrace>st i ** ps p ** one p\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1486
     i:[if_one e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1487
   \<lbrace>st e ** ps p ** one p\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1488
proof(unfold if_one_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1489
  fix l
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1490
  show " \<lbrace>st i \<and>* ps p \<and>* one p\<rbrace>  i :[ \<guillemotright> ((W0, l), W1, e) ]: l \<lbrace>st e \<and>* ps p \<and>* one p\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1491
  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1492
        intro tm.code_condI, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1493
    let ?j = "Suc i"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1494
    show "\<lbrace>one p \<and>* ps p \<and>* st i\<rbrace>  sg {TC i ((W0, ?j), W1, e)} \<lbrace>one p \<and>* ps p \<and>* st e\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1495
    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1496
      fix ft prog cs pc mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1497
      assume h: "(r \<and>* one p \<and>* ps p \<and>* st i \<and>* sg {TC i ((W0, ?j), W1, e)}) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1498
        (trset_of (ft, prog, cs, pc, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1499
      from h have k1: "prog $ i = Some ((W0, ?j), W1, e)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1500
        apply(rule_tac r = "r \<and>* one p \<and>* ps p" in codeD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1501
        by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1502
      from h have k2: "pc = p"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1503
        by(sep_drule psD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1504
      from h and this have k3: "mem $ pc = Some Oc"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1505
        apply(unfold one_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1506
        by(sep_drule memD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1507
      from h k1 k2 k3 have stp:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1508
        "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, e, pc, mem)" (is "?x = ?y")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1509
        apply(sep_drule stD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1510
        apply(unfold tm.run_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1511
        apply(auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1512
        thm fmap_eqI
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1513
        apply(rule fmap_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1514
        apply(simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1515
        apply(subgoal_tac "p \<in> fdom mem")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1516
        apply(simp add: insert_absorb)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1517
        apply(simp add: fdomIff)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1518
        by (metis the_lookup_fmap_upd the_lookup_fmap_upd_other)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1519
      from h k2 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1520
      have "(r \<and>* one p \<and>* ps p \<and>* st e \<and>* sg {TC i ((W0, ?j), W1, e)})  (trset_of ?x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1521
        apply(unfold stp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1522
        by(sep_drule st_upd, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1523
      thus "\<exists>k.(r \<and>* one p \<and>* ps p \<and>* st e \<and>* sg {TC i ((W0, ?j), W1, e)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1524
             (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1525
        apply (rule_tac x = 0 in exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1526
        by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1527
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1528
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1529
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1530
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1531
text {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1532
  The following problematic lemma is not provable now 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1533
  lemma hoare_self: "\<lbrace>p\<rbrace> i :[ap]: j \<lbrace>p\<rbrace>" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1534
  apply(simp add: tm.Hoare_gen_def, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1535
  apply(rule_tac x = 0 in exI, simp add: tm.run_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1536
  done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1537
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1538
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1539
lemma hoare_if_one_true_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1540
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1541
  shows
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1542
  "\<lbrace>st i ** ps p ** one q\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1543
     i:[if_one e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1544
   \<lbrace>st e ** ps p ** one q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1545
  by (unfold assms, rule hoare_if_one_true)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1546
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1547
lemma hoare_if_one_true1: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1548
  "\<lbrace>st i ** ps p ** one p\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1549
     i:[(if_one e; c)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1550
   \<lbrace>st e ** ps p ** one p\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1551
proof(unfold tassemble_to.simps, rule tm.code_exI, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1552
       simp add: sep_conj_ac tm.Hoare_gen_def, clarify)  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1553
  fix j' ft prog cs pos mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1554
  assume h: "(r \<and>* one p \<and>* ps p \<and>* st i \<and>* j' :[ c ]: j \<and>* i :[ if_one e ]: j') 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1555
    (trset_of (ft, prog, cs, pos, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1556
  from tm.frame_rule[OF hoare_if_one_true]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1557
  have "\<And> r. \<lbrace>st i \<and>* ps p \<and>* one p \<and>* r\<rbrace>  i :[ if_one e ]: j' \<lbrace>st e \<and>* ps p \<and>* one p \<and>* r\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1558
    by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1559
  from this[unfolded tm.Hoare_gen_def tassemble_to.simps, rule_format, of "j' :[ c ]: j"] h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1560
  have "\<exists> k. (r \<and>* one p \<and>* ps p \<and>* st e \<and>* i :[ if_one e ]: j' \<and>* j' :[ c ]: j)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1561
    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1562
    by(auto simp: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1563
  thus "\<exists>k. (r \<and>* one p \<and>* ps p \<and>* st e \<and>* j' :[ c ]: j \<and>* i :[ if_one e ]: j') 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1564
    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1565
    by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1566
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1567
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1568
lemma hoare_if_one_true1_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1569
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1570
  shows
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1571
  "\<lbrace>st i ** ps p ** one q\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1572
     i:[(if_one e; c)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1573
   \<lbrace>st e ** ps p ** one q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1574
  by (unfold assms, rule hoare_if_one_true1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1575
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1576
lemma hoare_if_one_false: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1577
  "\<lbrace>st i ** ps p ** zero p\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1578
       i:[if_one e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1579
   \<lbrace>st j ** ps p ** zero p\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1580
proof(unfold if_one_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1581
  show "\<lbrace>st i \<and>* ps p \<and>* zero p\<rbrace>  i :[ (\<guillemotright> ((W0, j), W1, e)) ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1582
        \<lbrace>st j \<and>* ps p \<and>* zero p\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1583
  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1584
        intro tm.code_condI, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1585
    let ?j = "Suc i"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1586
    show "\<lbrace>ps p \<and>* st i \<and>* zero p\<rbrace>  sg {TC i ((W0, ?j), W1, e)} \<lbrace>ps p \<and>*  zero p \<and>* st ?j \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1587
    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1588
      fix ft prog cs pc mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1589
      assume h: "(r \<and>* ps p \<and>* st i \<and>* zero p \<and>* sg {TC i ((W0, ?j), W1, e)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1590
        (trset_of (ft, prog, cs, pc, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1591
      from h have k1: "prog $ i = Some ((W0, ?j), W1, e)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1592
        apply(rule_tac r = "r \<and>* zero p \<and>* ps p" in codeD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1593
        by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1594
      from h have k2: "pc = p"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1595
        by(sep_drule psD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1596
      from h and this have k3: "mem $ pc = Some Bk"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1597
        apply(unfold zero_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1598
        by(sep_drule memD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1599
      from h k1 k2 k3 have stp:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1600
        "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, ?j, pc, mem)" (is "?x = ?y")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1601
        apply(sep_drule stD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1602
        apply(unfold tm.run_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1603
        apply(auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1604
        apply(rule fmap_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1605
        apply(simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1606
        apply(subgoal_tac "p \<in> fdom mem")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1607
        apply(simp add: insert_absorb)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1608
        apply(simp add: fdomIff)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1609
        by (metis the_lookup_fmap_upd the_lookup_fmap_upd_other)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1610
      from h k2 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1611
      have "(r \<and>* zero p \<and>* ps p \<and>* st ?j \<and>* sg {TC i ((W0, ?j), W1, e)})  (trset_of ?x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1612
        apply (unfold stp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1613
        by (sep_drule st_upd[where i''="?j"], auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1614
      thus "\<exists>k. (r \<and>* ps p \<and>* zero p \<and>* st ?j \<and>*  sg {TC i ((W0, ?j), W1, e)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1615
             (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1616
        by(auto simp: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1617
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1618
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1619
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1620
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1621
lemma hoare_if_one_false_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1622
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1623
  shows "\<lbrace>st i ** ps p ** zero q\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1624
             i:[if_one e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1625
         \<lbrace>st j ** ps p ** zero q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1626
  by (unfold assms, rule hoare_if_one_false)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1627
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1628
definition "if_zero e = (TL exit . \<guillemotright>((W0, e), (W1, exit)); TLabel exit)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1629
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1630
lemma hoare_if_zero_true: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1631
  "\<lbrace>st i ** ps p ** zero p\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1632
     i:[if_zero e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1633
   \<lbrace>st e ** ps p ** zero p\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1634
proof(unfold if_zero_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1635
  fix l
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1636
  show "\<lbrace>st i \<and>* ps p \<and>* zero p\<rbrace>  i :[ \<guillemotright> ((W0, e), W1, l) ]: l \<lbrace>st e \<and>* ps p \<and>* zero p\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1637
  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1638
        intro tm.code_condI, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1639
    let ?j = "Suc i"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1640
    show "\<lbrace>ps p \<and>* st i \<and>* zero p\<rbrace>  sg {TC i ((W0, e), W1, ?j)} \<lbrace>ps p \<and>* st e \<and>* zero p\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1641
    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1642
      fix ft prog cs pc mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1643
      assume h: "(r \<and>* ps p \<and>* st i \<and>* zero p \<and>* sg {TC i ((W0, e), W1, ?j)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1644
        (trset_of (ft, prog, cs, pc, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1645
      from h have k1: "prog $ i = Some ((W0, e), W1, ?j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1646
        apply(rule_tac r = "r \<and>* zero p \<and>* ps p" in codeD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1647
        by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1648
      from h have k2: "pc = p"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1649
        by(sep_drule psD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1650
      from h and this have k3: "mem $ pc = Some Bk"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1651
        apply(unfold zero_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1652
        by(sep_drule memD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1653
      from h k1 k2 k3 have stp:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1654
        "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, e, pc, mem)" (is "?x = ?y")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1655
        apply(sep_drule stD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1656
        apply(unfold tm.run_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1657
        apply(auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1658
        apply(rule fmap_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1659
        apply(simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1660
        apply(subgoal_tac "p \<in> fdom mem")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1661
        apply(simp add: insert_absorb)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1662
        apply(simp add: fdomIff)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1663
        by (metis the_lookup_fmap_upd the_lookup_fmap_upd_other)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1664
      from h k2 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1665
      have "(r \<and>* zero p \<and>* ps p \<and>* st e \<and>* sg {TC i ((W0, e), W1, ?j)})  (trset_of ?x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1666
        apply(unfold stp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1667
        by(sep_drule st_upd, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1668
      thus "\<exists>k. (r \<and>* ps p \<and>* st e \<and>* zero p \<and>* sg {TC i ((W0, e), W1, ?j)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1669
             (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1670
        by(auto simp: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1671
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1672
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1673
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1674
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1675
lemma hoare_if_zero_true_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1676
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1677
  shows
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1678
  "\<lbrace>st i ** ps p ** zero q\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1679
     i:[if_zero e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1680
   \<lbrace>st e ** ps p ** zero q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1681
  by (unfold assms, rule hoare_if_zero_true)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1682
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1683
lemma hoare_if_zero_true1: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1684
  "\<lbrace>st i ** ps p ** zero p\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1685
     i:[(if_zero e; c)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1686
   \<lbrace>st e ** ps p ** zero p\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1687
 proof(unfold tassemble_to.simps, rule tm.code_exI, simp add: sep_conj_ac 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1688
        tm.Hoare_gen_def, clarify)  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1689
  fix j' ft prog cs pos mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1690
  assume h: "(r \<and>* ps p \<and>* st i \<and>* zero p \<and>* j' :[ c ]: j \<and>* i :[ if_zero e ]: j') 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1691
    (trset_of (ft, prog, cs, pos, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1692
  from tm.frame_rule[OF hoare_if_zero_true]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1693
  have "\<And> r. \<lbrace>st i \<and>* ps p \<and>* zero p \<and>* r\<rbrace>  i :[ if_zero e ]: j' \<lbrace>st e \<and>* ps p \<and>* zero p \<and>* r\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1694
    by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1695
  from this[unfolded tm.Hoare_gen_def tassemble_to.simps, rule_format, of "j' :[ c ]: j"] h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1696
  have "\<exists> k. (r \<and>* zero p \<and>* ps p \<and>* st e \<and>* i :[ if_zero e ]: j' \<and>* j' :[ c ]: j)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1697
    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1698
    by(auto simp: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1699
  thus "\<exists>k. (r \<and>* ps p \<and>* st e \<and>* zero p \<and>* j' :[ c ]: j \<and>* i :[ if_zero e ]: j')  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1700
    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1701
    by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1702
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1703
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1704
lemma hoare_if_zero_true1_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1705
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1706
  shows
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1707
  "\<lbrace>st i ** ps p ** zero q\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1708
     i:[(if_zero e; c)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1709
   \<lbrace>st e ** ps p ** zero q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1710
  by (unfold assms, rule hoare_if_zero_true1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1711
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1712
lemma hoare_if_zero_false: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1713
  "\<lbrace>st i ** ps p ** tm p Oc\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1714
     i:[if_zero e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1715
   \<lbrace>st j ** ps p ** tm p Oc\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1716
proof(unfold if_zero_def, intro t_hoare_local, rule t_hoare_label_last, simp, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1717
  fix l
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1718
  show "\<lbrace>st i \<and>* ps p \<and>* tm p Oc\<rbrace>  i :[ \<guillemotright> ((W0, e), W1, l) ]: l
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1719
        \<lbrace>st l \<and>* ps p \<and>* tm p Oc\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1720
  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1721
      intro tm.code_condI, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1722
    let ?j = "Suc i"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1723
    show "\<lbrace>ps p \<and>* st i \<and>* tm p Oc\<rbrace>  sg {TC i ((W0, e), W1, ?j)} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1724
          \<lbrace>ps p \<and>* st ?j \<and>* tm p Oc\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1725
    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1726
      fix ft prog cs pc mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1727
      assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p Oc \<and>* sg {TC i ((W0, e), W1, ?j)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1728
        (trset_of (ft, prog, cs, pc, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1729
      from h have k1: "prog $ i = Some ((W0, e), W1, ?j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1730
        apply(rule_tac r = "r \<and>* tm p Oc \<and>* ps p" in codeD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1731
        by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1732
      from h have k2: "pc = p"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1733
        by(sep_drule psD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1734
      from h and this have k3: "mem $ pc = Some Oc"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1735
        by(sep_drule memD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1736
      from h k1 k2 k3 have stp:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1737
        "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, ?j, pc, mem)" (is "?x = ?y")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1738
        apply(sep_drule stD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1739
        apply(unfold tm.run_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1740
        apply(auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1741
        apply(rule fmap_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1742
        apply(simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1743
        apply(subgoal_tac "p \<in> fdom mem")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1744
        apply(simp add: insert_absorb)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1745
        apply(simp add: fdomIff)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1746
        by (metis the_lookup_fmap_upd the_lookup_fmap_upd_other)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1747
      from h k2 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1748
      have "(r \<and>* tm p Oc \<and>* ps p \<and>* st ?j \<and>* sg {TC i ((W0, e), W1, ?j)})  (trset_of ?x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1749
        apply(unfold stp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1750
        by(sep_drule st_upd, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1751
      thus "\<exists>k. (r \<and>* ps p \<and>* st ?j \<and>* tm p Oc \<and>* sg {TC i ((W0, e), W1, ?j)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1752
             (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1753
        by(auto simp: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1754
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1755
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1756
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1757
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1758
lemma hoare_if_zero_false_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1759
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1760
  shows
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1761
  "\<lbrace>st i ** ps p ** tm q Oc\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1762
     i:[if_zero e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1763
   \<lbrace>st j ** ps p ** tm q Oc\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1764
  by (unfold assms, rule hoare_if_zero_false)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1765
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1766
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1767
definition "jmp e = \<guillemotright>((W0, e), (W1, e))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1768
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1769
lemma hoare_jmp: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1770
  "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace>  i:[jmp e]:j \<lbrace>st e \<and>* ps p \<and>* tm p v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1771
proof(unfold jmp_def tm.Hoare_gen_def tassemble_to.simps sep_conj_ac, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1772
  fix ft prog cs pos mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1773
  assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* <(j = i + 1)> \<and>* sg {TC i ((W0, e), W1, e)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1774
    (trset_of (ft, prog, cs, pos, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1775
  from h have k1: "prog $ i = Some ((W0, e), W1, e)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1776
    apply(rule_tac r = "r \<and>* <(j = i + 1)> \<and>* tm p v \<and>* ps p" in codeD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1777
    by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1778
  from h have k2: "p = pos"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1779
    by(sep_drule psD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1780
  from h k2 have k3: "mem $ pos = Some v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1781
    by(sep_drule memD, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1782
  from h k1 k2 k3 have 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1783
    stp: "tm.run 1 (ft, prog, cs, pos, mem) = (ft, prog, e, pos, mem)" (is "?x = ?y")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1784
    apply(sep_drule stD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1785
    apply(unfold tm.run_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1786
    apply(cases "mem $ pos")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1787
    apply(simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1788
    apply(cases v)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1789
    apply(auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1790
    apply(rule fmap_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1791
    apply(simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1792
    apply(subgoal_tac "pos \<in> fdom mem")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1793
    apply(simp add: insert_absorb)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1794
    apply(simp add: fdomIff)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1795
    apply(metis the_lookup_fmap_upd the_lookup_fmap_upd_other)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1796
    apply(rule fmap_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1797
    apply(simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1798
    apply(subgoal_tac "pos \<in> fdom mem")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1799
    apply(simp add: insert_absorb)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1800
    apply(simp add: fdomIff)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1801
    apply(metis the_lookup_fmap_upd the_lookup_fmap_upd_other)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1802
    done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1803
  from h k2 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1804
  have "(r \<and>* ps p \<and>* st e \<and>* tm p v \<and>* <(j = i + 1)> \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1805
           sg {TC i ((W0, e), W1, e)}) (trset_of ?x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1806
    apply(unfold stp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1807
    by(sep_drule st_upd, simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1808
  thus "\<exists> k. (r \<and>* ps p \<and>* st e \<and>* tm p v \<and>* <(j = i + 1)> \<and>* sg {TC i ((W0, e), W1, e)})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1809
             (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1810
    apply (rule_tac x = 0 in exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1811
    by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1812
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1813
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1814
lemma hoare_jmp_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1815
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1816
  shows "\<lbrace>st i \<and>* ps p \<and>* tm q v\<rbrace>  i:[jmp e]:j \<lbrace>st e \<and>* ps p \<and>* tm q v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1817
  by (unfold assms, rule hoare_jmp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1818
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1819
lemma hoare_jmp1: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1820
  "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1821
     i:[(jmp e; c)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1822
   \<lbrace>st e \<and>* ps p \<and>* tm p v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1823
proof(unfold  tassemble_to.simps, rule tm.code_exI, simp 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1824
              add: sep_conj_ac tm.Hoare_gen_def, clarify)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1825
  fix j' ft prog cs pos mem r
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1826
  assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* j' :[ c ]: j \<and>* i :[ jmp e ]: j') 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1827
    (trset_of (ft, prog, cs, pos, mem))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1828
  from tm.frame_rule[OF hoare_jmp]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1829
  have "\<And> r. \<lbrace>st i \<and>* ps p \<and>* tm p v \<and>* r\<rbrace>  i :[ jmp e ]: j' \<lbrace>st e \<and>* ps p \<and>* tm p v \<and>* r\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1830
    by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1831
  from this[unfolded tm.Hoare_gen_def tassemble_to.simps, rule_format, of "j' :[ c ]: j"] h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1832
  have "\<exists> k. (r \<and>* tm p v \<and>* ps p \<and>* st e \<and>* i :[ jmp e ]: j' \<and>* j' :[ c ]: j)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1833
    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1834
    by(auto simp: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1835
  thus "\<exists>k. (r \<and>* ps p \<and>* st e \<and>* tm p v \<and>* j' :[ c ]: j \<and>* i :[ jmp e ]: j')  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1836
    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1837
    by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1838
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1839
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1840
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1841
lemma hoare_jmp1_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1842
  assumes "p = q"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1843
  shows "\<lbrace>st i \<and>* ps p \<and>* tm q v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1844
            i:[(jmp e; c)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1845
         \<lbrace>st e \<and>* ps p \<and>* tm q v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1846
  by (unfold assms, rule hoare_jmp1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1847
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1848
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1849
lemma condI: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1850
  assumes h1: b
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1851
  and h2: "b \<Longrightarrow> p s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1852
  shows "(<b> \<and>* p) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1853
  by (metis (full_types) cond_true_eq1 h1 h2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1854
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1855
lemma condE:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1856
  assumes "(<b> \<and>* p) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1857
  obtains "b" and "p s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1858
proof(atomize_elim)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1859
  from condD[OF assms]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1860
  show "b \<and> p s" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1861
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1862
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1863
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1864
section {* Tactics *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1865
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1866
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1867
  val trace_step = Attrib.setup_config_bool @{binding trace_step} (K false)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1868
  val trace_fwd = Attrib.setup_config_bool @{binding trace_fwd} (K false)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1869
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1870
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1871
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1872
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1873
  val tracing  = (fn ctxt => fn str =>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1874
                   if (Config.get ctxt trace_step) then tracing str else ())
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1875
  fun not_pred p = fn s => not (p s)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1876
  fun break_sep_conj (Const (@{const_name sep_conj},_) $ t1 $ t2 $ _) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1877
         (break_sep_conj t1) @ (break_sep_conj t2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1878
    | break_sep_conj (Const (@{const_name sep_conj},_) $ t1 $ t2) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1879
            (break_sep_conj t1) @ (break_sep_conj t2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1880
                   (* dig through eta exanded terms: *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1881
    | break_sep_conj (Abs (_, _, t $ Bound 0)) = break_sep_conj t
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1882
    | break_sep_conj t = [t];
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1883
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1884
  val empty_env = (Vartab.empty, Vartab.empty)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1885
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1886
  fun match_env ctxt pat trm env = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1887
            Pattern.match (ctxt |> Proof_Context.theory_of) (pat, trm) env
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1888
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1889
  fun match ctxt pat trm = match_env ctxt pat trm empty_env;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1890
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1891
  val inst = Envir.subst_term;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1892
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1893
  fun term_of_thm thm = thm |>  prop_of |> HOLogic.dest_Trueprop
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1894
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1895
  fun get_cmd ctxt code = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1896
      let val pat = term_of @{cpat "_:[(?cmd)]:_"}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1897
          val pat1 = term_of @{cpat "?cmd::tpg"}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1898
          val env = match ctxt pat code
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1899
      in inst env pat1 end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1900
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1901
  fun is_seq_term (Const (@{const_name TSeq}, _) $ _ $ _) = true
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1902
    | is_seq_term _ = false
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1903
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1904
  fun get_hcmd  (Const (@{const_name TSeq}, _) $ hcmd $ _) = hcmd
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1905
    | get_hcmd hcmd = hcmd
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1906
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1907
  fun last [a]  = a |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1908
      last (a::b) = last b
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1909
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1910
  fun but_last [a] = [] |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1911
      but_last (a::b) = a::(but_last b)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1912
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1913
  fun foldr f [] = (fn x => x) |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1914
      foldr f (x :: xs) = (f x) o  (foldr f xs)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1915
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1916
  fun concat [] = [] |
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1917
      concat (x :: xs) = x @ concat xs
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1918
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1919
  fun match_any ctxt pats tm = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1920
              fold 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1921
                 (fn pat => fn b => (b orelse Pattern.matches 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1922
                          (ctxt |> Proof_Context.theory_of) (pat, tm))) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1923
                 pats false
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1924
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1925
  fun is_ps_term (Const (@{const_name ps}, _) $ _) = true
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1926
    | is_ps_term _ = false
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1927
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1928
  fun string_of_term ctxt t = t |> Syntax.pretty_term ctxt |> Pretty.str_of
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1929
  fun string_of_cterm ctxt ct = ct |> term_of |> string_of_term ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1930
  fun pterm ctxt t =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1931
          t |> string_of_term ctxt |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1932
  fun pcterm ctxt ct = ct |> string_of_cterm ctxt |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1933
  fun string_for_term ctxt t =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1934
       Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1935
                   (print_mode_value ())) (Syntax.string_of_term ctxt) t
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1936
         |> String.translate (fn c => if Char.isPrint c then str c else "")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1937
         |> Sledgehammer_Util.simplify_spaces  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1938
  fun string_for_cterm ctxt ct = ct |> term_of |> string_for_term ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1939
  fun attemp tac = fn i => fn st => (tac i st) handle exn => Seq.empty
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1940
  fun try_tac tac = fn i => fn st => (tac i st) handle exn => (Seq.single st)       
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1941
 (* aux end *) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1942
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1943
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1944
ML {* (* Functions specific to Hoare triples *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1945
  fun get_pre ctxt t = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1946
    let val pat = term_of @{cpat "\<lbrace>?P\<rbrace> ?c \<lbrace>?Q\<rbrace>"} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1947
        val env = match ctxt pat t 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1948
    in inst env (term_of @{cpat "?P::tresource set \<Rightarrow> bool"}) end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1949
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1950
  fun can_process ctxt t = ((get_pre ctxt t; true) handle _ => false)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1951
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1952
  fun get_post ctxt t = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1953
    let val pat = term_of @{cpat "\<lbrace>?P\<rbrace> ?c \<lbrace>?Q\<rbrace>"} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1954
        val env = match ctxt pat t 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1955
    in inst env (term_of @{cpat "?Q::tresource set \<Rightarrow> bool"}) end;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1956
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1957
  fun get_mid ctxt t = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1958
    let val pat = term_of @{cpat "\<lbrace>?P\<rbrace> ?c \<lbrace>?Q\<rbrace>"} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1959
        val env = match ctxt pat t 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1960
    in inst env (term_of @{cpat "?c::tresource set \<Rightarrow> bool"}) end;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1961
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1962
  fun is_pc_term (Const (@{const_name st}, _) $ _) = true
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1963
    | is_pc_term _ = false
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1964
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1965
  fun mk_pc_term x =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1966
     Const (@{const_name st}, @{typ "nat \<Rightarrow> tresource set \<Rightarrow> bool"}) $ Free (x, @{typ "nat"})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1967
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1968
  val sconj_term = term_of @{cterm "sep_conj::tassert \<Rightarrow> tassert \<Rightarrow> tassert"}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1969
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1970
  fun mk_ps_term x =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1971
     Const (@{const_name ps}, @{typ "int \<Rightarrow> tresource set \<Rightarrow> bool"}) $ Free (x, @{typ "int"})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1972
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1973
  fun atomic tac  = ((SOLVED' tac) ORELSE' (K all_tac))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1974
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1975
  fun map_simpset f = Context.proof_map (Simplifier.map_ss f)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1976
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1977
  fun pure_sep_conj_ac_tac ctxt = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1978
         (auto_tac (map_simpset (fn ss => ss addsimps @{thms sep_conj_ac}) ctxt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1979
          |> SELECT_GOAL)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1980
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1981
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1982
  fun potential_facts ctxt prop = Facts.could_unify (Proof_Context.facts_of ctxt) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1983
                                       ((Term.strip_all_body prop) |> Logic.strip_imp_concl);
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1984
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1985
  fun some_fact_tac ctxt = SUBGOAL (fn (goal, i) => 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1986
                                      (Method.insert_tac (potential_facts ctxt goal) i) THEN
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1987
                                      (pure_sep_conj_ac_tac ctxt i));
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1988
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1989
  fun sep_conj_ac_tac ctxt = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1990
     (SOLVED' (auto_tac (ctxt |> map_simpset (fn ss => ss addsimps @{thms sep_conj_ac}))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1991
       |> SELECT_GOAL)) ORELSE' (atomic (some_fact_tac ctxt))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1992
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1993
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1994
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1995
type HoareTriple = {
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1996
  binding: binding,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1997
  can_process: Proof.context -> term -> bool,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1998
  get_pre: Proof.context -> term -> term,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1999
  get_mid: Proof.context -> term -> term,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2000
  get_post: Proof.context -> term -> term,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2001
  is_pc_term: term -> bool,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2002
  mk_pc_term: string -> term,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2003
  sconj_term: term,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2004
  sep_conj_ac_tac: Proof.context -> int -> tactic,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2005
  hoare_seq1: thm,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2006
  hoare_seq2: thm,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2007
  pre_stren: thm,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2008
  post_weaken: thm,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2009
  frame_rule: thm
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2010
}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2011
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2012
  val tm_triple = {binding = @{binding "tm_triple"}, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2013
                   can_process = can_process,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2014
                   get_pre = get_pre,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2015
                   get_mid = get_mid,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2016
                   get_post = get_post,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2017
                   is_pc_term = is_pc_term,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2018
                   mk_pc_term = mk_pc_term,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2019
                   sconj_term = sconj_term,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2020
                   sep_conj_ac_tac = sep_conj_ac_tac,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2021
                   hoare_seq1 = @{thm t_hoare_seq1},
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2022
                   hoare_seq2 = @{thm t_hoare_seq2},
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2023
                   pre_stren = @{thm tm.pre_stren},
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2024
                   post_weaken = @{thm tm.post_weaken},
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2025
                   frame_rule = @{thm tm.frame_rule}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2026
                  }:HoareTriple
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2027
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2028
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2029
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2030
  val _ = data_slot "HoareTriples" "HoareTriple list" "[]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2031
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2032
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2033
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2034
  val _ = HoareTriples_store [tm_triple]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2035
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2036
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2037
ML {* (* aux1 functions *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2038
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2039
fun focus_params t ctxt =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2040
  let
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2041
    val (xs, Ts) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2042
      split_list (Term.variant_frees t (Term.strip_all_vars t));  (*as they are printed :-*)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2043
    (* val (xs', ctxt') = variant_fixes xs ctxt; *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2044
    (* val ps = xs' ~~ Ts; *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2045
    val ps = xs ~~ Ts
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2046
    val (_, ctxt'') = ctxt |> Variable.add_fixes xs
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2047
  in ((xs, ps), ctxt'') end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2048
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2049
fun focus_concl ctxt t =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2050
  let
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2051
    val ((xs, ps), ctxt') = focus_params t ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2052
    val t' = Term.subst_bounds (rev (map Free ps), Term.strip_all_body t);
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2053
  in (t' |> Logic.strip_imp_concl, ctxt') end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2054
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2055
  fun get_concl ctxt (i, state) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2056
              nth (Thm.prems_of state) (i - 1) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2057
                            |> focus_concl ctxt |> (fn (x, _) => x |> HOLogic.dest_Trueprop)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2058
 (* aux1 end *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2059
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2060
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2061
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2062
  fun indexing xs = upto (0, length xs - 1) ~~ xs
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2063
  fun select_idxs idxs ps = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2064
      map_index (fn (i, e) => if (member (op =) idxs i) then [e] else []) ps |> flat
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2065
  fun select_out_idxs idxs ps = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2066
      map_index (fn (i, e) => if (member (op =) idxs i) then [] else [e]) ps |> flat
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2067
  fun match_pres ctxt mf env ps qs = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2068
      let  fun sel_match mf env [] qs = [(env, [])]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2069
             | sel_match mf env (p::ps) qs = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2070
                  let val pm = map (fn (i, q) => [(i, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2071
                                      let val _ = tracing ctxt "Matching:"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2072
                                          val _ = [p, q] |>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2073
                                            (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2074
                                          val r = mf p q env 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2075
                                      in r end)]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2076
                                      handle _ => (
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2077
                                      let val _ = tracing ctxt "Failed matching:"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2078
                                          val _ = [p, q] |>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2079
                                            (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2080
                                      in [] end)) qs |> flat
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2081
                      val r = pm |> map (fn (i, env') => 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2082
                                let val qs' = filter_out (fn (j, q) => j = i) qs
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2083
                                in  sel_match mf env' ps qs' |> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2084
                                      map (fn (env'', idxs) => (env'', i::idxs)) end) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2085
                        |> flat
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2086
            in r end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2087
   in sel_match mf env ps (indexing qs) end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2088
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2089
  fun provable tac ctxt goal = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2090
          let 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2091
              val _ = tracing ctxt "Provable trying to prove:"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2092
              val _ = [goal] |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2093
          in
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2094
             (Goal.prove ctxt [] [] goal (fn {context, ...} => tac context 1); true)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2095
                        handle exn => false
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2096
          end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2097
  fun make_sense tac ctxt thm_assms env  = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2098
                thm_assms |>  map (inst env) |> forall (provable tac ctxt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2099
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2100
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2101
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2102
  fun triple_for ctxt goal = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2103
    filter (fn trpl => (#can_process trpl) ctxt goal) (HoareTriples.get (Proof_Context.theory_of ctxt)) |> hd
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2104
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2105
  fun step_terms_for thm goal ctxt = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2106
    let
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2107
       val _ = tracing ctxt "This is the new version of step_terms_for!"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2108
       val _ = tracing ctxt "Tring to find triple processor: TP"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2109
       val TP = triple_for ctxt goal
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2110
       val _ = #binding TP |> Binding.name_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2111
       fun mk_sep_conj tms = foldr (fn tm => fn rtm => 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2112
              ((#sconj_term TP)$tm$rtm)) (but_last tms) (last tms)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2113
       val thm_concl = thm |> prop_of 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2114
                 |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2115
       val thm_assms = thm |> prop_of 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2116
           |> Logic.strip_imp_prems 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2117
       val cmd_pat = thm_concl |> #get_mid TP ctxt |> get_cmd ctxt 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2118
       val cmd = goal |> #get_mid TP ctxt |> get_cmd ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2119
       val _ = tracing ctxt "matching command ... "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2120
       val _ = tracing ctxt "cmd_pat = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2121
       val _ = pterm ctxt cmd_pat
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2122
       val (hcmd, env1, is_last) =  (cmd, match ctxt cmd_pat cmd, true)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2123
             handle exn => (cmd |> get_hcmd, match ctxt cmd_pat (cmd |> get_hcmd), false)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2124
       val _ = tracing ctxt "hcmd ="
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2125
       val _ = pterm ctxt hcmd
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2126
       val _ = tracing ctxt "match command succeed! "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2127
       val _ = tracing ctxt "pres ="
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2128
       val pres = goal |> #get_pre TP ctxt |> break_sep_conj 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2129
       val _ = pres |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2130
       val _ = tracing ctxt "pre_pats ="
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2131
       val pre_pats = thm_concl |> #get_pre TP ctxt |> inst env1 |> break_sep_conj
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2132
       val _ = pre_pats |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2133
       val _ = tracing ctxt "post_pats ="
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2134
       val post_pats = thm_concl |> #get_post TP ctxt |> inst env1 |> break_sep_conj
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2135
       val _ = post_pats |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2136
       val _ = tracing ctxt "Calculating sols"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2137
       val sols = match_pres ctxt (match_env ctxt) env1 pre_pats pres 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2138
       val _ = tracing ctxt "End calculating sols, sols ="
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2139
       val _ = tracing ctxt (@{make_string} sols)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2140
       val _ = tracing ctxt "Calulating env2 and idxs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2141
       val (env2, idxs) = filter (fn (env, idxs) => make_sense (#sep_conj_ac_tac TP) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2142
                             ctxt thm_assms env) sols |> hd
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2143
       val _ = tracing ctxt "End calculating env2 and idxs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2144
       val _ = tracing ctxt "mterms ="
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2145
       val mterms = select_idxs idxs pres |> map (inst env2) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2146
       val _ = mterms |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2147
       val _ = tracing ctxt "nmterms = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2148
       val nmterms = select_out_idxs idxs pres |> map (inst env2) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2149
       val _ = nmterms |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2150
       val pre_cond = pre_pats |> map (inst env2) |> mk_sep_conj
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2151
       val post_cond = post_pats |> map (inst env2) |> mk_sep_conj 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2152
       val post_cond_npc  = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2153
               post_cond |> break_sep_conj |> filter (not_pred (#is_pc_term TP)) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2154
               |> (fn x => x @ nmterms) |> mk_sep_conj |> cterm_of (Proof_Context.theory_of ctxt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2155
       fun mk_frame cond rest  = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2156
             if rest = [] then cond else ((#sconj_term TP)$ cond) $ (mk_sep_conj rest)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2157
       val pre_cond_frame = mk_frame pre_cond nmterms |> cterm_of (Proof_Context.theory_of ctxt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2158
       fun post_cond_frame j' = post_cond |> break_sep_conj |> filter (not_pred (#is_pc_term TP)) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2159
               |> (fn x => [#mk_pc_term TP j']@x) |> mk_sep_conj
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2160
               |> (fn x => mk_frame x nmterms)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2161
               |> cterm_of (Proof_Context.theory_of ctxt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2162
       val need_frame = (nmterms <> [])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2163
    in 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2164
         (post_cond_npc,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2165
          pre_cond_frame, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2166
          post_cond_frame, need_frame, is_last)       
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2167
    end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2168
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2169
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2170
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2171
  fun step_tac ctxt thm i state = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2172
     let  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2173
       val _ = tracing ctxt "This is the new version of step_tac"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2174
       val (goal, ctxt) = nth (Thm.prems_of state) (i - 1) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2175
                  |> focus_concl ctxt 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2176
                  |> (apfst HOLogic.dest_Trueprop)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2177
       val _ = tracing ctxt "step_tac: goal = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2178
       val _ = goal |> pterm ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2179
       val _ = tracing ctxt "Start to calculate intermediate terms ... "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2180
       val (post_cond_npc, pre_cond_frame, post_cond_frame, need_frame, is_last) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2181
                        = step_terms_for thm goal ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2182
       val _ = tracing ctxt "Tring to find triple processor: TP"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2183
       val TP = triple_for ctxt goal
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2184
       val _ = #binding TP |> Binding.name_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2185
       fun mk_sep_conj tms = foldr (fn tm => fn rtm => 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2186
              ((#sconj_term TP)$tm$rtm)) (but_last tms) (last tms)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2187
       val _ = tracing ctxt "Calculate intermediate terms finished! "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2188
       val post_cond_npc_str = post_cond_npc |> string_for_cterm ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2189
       val pre_cond_frame_str = pre_cond_frame |> string_for_cterm ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2190
       val _ = tracing ctxt "step_tac: post_cond_npc = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2191
       val _ = post_cond_npc |> pcterm ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2192
       val _ = tracing ctxt "step_tac: pre_cond_frame = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2193
       val _ = pre_cond_frame |> pcterm ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2194
       fun tac1 i state = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2195
             if is_last then (K all_tac) i state else
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2196
              res_inst_tac ctxt [(("q", 0), post_cond_npc_str)] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2197
                                          (#hoare_seq1 TP) i state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2198
       fun tac2 i state = res_inst_tac ctxt [(("p", 0), pre_cond_frame_str)] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2199
                                          (#pre_stren TP) i state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2200
       fun foc_tac post_cond_frame ctxt i state  =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2201
           let
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2202
               val goal = get_concl ctxt (i, state)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2203
               val pc_term = goal |> #get_post TP ctxt |> break_sep_conj 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2204
                                |> filter (#is_pc_term TP) |> hd
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2205
               val (_$Free(j', _)) = pc_term
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2206
               val psd = post_cond_frame j'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2207
               val str_psd = psd |> string_for_cterm ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2208
               val _ = tracing ctxt "foc_tac: psd = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2209
               val _ = psd |> pcterm ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2210
           in 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2211
               res_inst_tac ctxt [(("q", 0), str_psd)] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2212
                                          (#post_weaken TP) i state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2213
           end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2214
     val frame_tac = if need_frame then (rtac (#frame_rule TP)) else (K all_tac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2215
     val print_tac = if (Config.get ctxt trace_step) then Tactical.print_tac else (K all_tac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2216
     val tac = (tac1 THEN' (K (print_tac "tac1 success"))) THEN' 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2217
               (tac2 THEN' (K (print_tac "tac2 success"))) THEN' 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2218
               ((foc_tac post_cond_frame ctxt) THEN' (K (print_tac "foc_tac success"))) THEN' 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2219
               (frame_tac  THEN' (K (print_tac "frame_tac success"))) THEN' 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2220
               (((rtac thm) THEN_ALL_NEW (#sep_conj_ac_tac TP ctxt)) THEN' (K (print_tac "rtac thm success"))) THEN' 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2221
               (K (ALLGOALS (atomic (#sep_conj_ac_tac TP ctxt)))) THEN'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2222
               (* (#sep_conj_ac_tac TP ctxt) THEN' (#sep_conj_ac_tac TP ctxt) THEN'  *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2223
               (K prune_params_tac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2224
   in 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2225
        tac i state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2226
   end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2227
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2228
  fun unfold_cell_tac ctxt = (Local_Defs.unfold_tac ctxt @{thms one_def zero_def})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2229
  fun fold_cell_tac ctxt = (Local_Defs.fold_tac ctxt @{thms one_def zero_def})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2230
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2231
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2232
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2233
  fun sg_step_tac thms ctxt =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2234
     let val sg_step_tac' =  (map (fn thm  => attemp (step_tac ctxt thm)) thms)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2235
                               (* @ [attemp (goto_tac ctxt)]  *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2236
                              |> FIRST'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2237
         val sg_step_tac'' = (K (unfold_cell_tac ctxt)) THEN' sg_step_tac' THEN' (K (fold_cell_tac ctxt))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2238
     in
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2239
         sg_step_tac' ORELSE' sg_step_tac''
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2240
     end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2241
  fun steps_tac thms ctxt i = REPEAT (sg_step_tac thms ctxt i) THEN (prune_params_tac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2242
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2243
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2244
method_setup hstep = {* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2245
  Attrib.thms >> (fn thms => fn ctxt =>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2246
                    (SIMPLE_METHOD' (fn i => 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2247
                       sg_step_tac (thms@(StepRules.get ctxt)) ctxt i)))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2248
  *} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2249
  "One step symbolic execution using step theorems."
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2250
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2251
method_setup hsteps = {* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2252
  Attrib.thms >> (fn thms => fn ctxt =>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2253
                    (SIMPLE_METHOD' (fn i => 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2254
                       steps_tac (thms@(StepRules.get ctxt)) ctxt i)))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2255
  *} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2256
  "Sequential symbolic execution using step theorems."
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2257
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2258
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2259
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2260
  fun goto_tac ctxt thm i state = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2261
     let  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2262
       val (goal, ctxt) = nth (Thm.prems_of state) (i - 1) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2263
                             |> focus_concl ctxt |> (apfst HOLogic.dest_Trueprop)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2264
       val _ = tracing ctxt "goto_tac: goal = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2265
       val _ = goal |> string_of_term ctxt |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2266
       val (post_cond_npc, pre_cond_frame, post_cond_frame, need_frame, is_last) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2267
                        = step_terms_for thm goal ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2268
       val _ = tracing ctxt "Tring to find triple processor: TP"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2269
       val TP = triple_for ctxt goal
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2270
       val _ = #binding TP |> Binding.name_of |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2271
       val _ = tracing ctxt "aaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaa"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2272
       val post_cond_npc_str = post_cond_npc |> string_for_cterm ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2273
       val pre_cond_frame_str = pre_cond_frame |> string_for_cterm ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2274
       val _ = tracing ctxt "goto_tac: post_cond_npc = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2275
       val _ = post_cond_npc_str |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2276
       val _ = tracing ctxt "goto_tac: pre_cond_frame = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2277
       val _ = pre_cond_frame_str |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2278
       fun tac1 i state = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2279
             if is_last then (K all_tac) i state else
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2280
              res_inst_tac ctxt [] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2281
                                          (#hoare_seq2 TP) i state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2282
       fun tac2 i state = res_inst_tac ctxt [(("p", 0), pre_cond_frame_str)] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2283
                                          (#pre_stren TP) i state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2284
       fun foc_tac post_cond_frame ctxt i state  =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2285
           let
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2286
               val goal = get_concl ctxt (i, state)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2287
               val pc_term = goal |> #get_post TP ctxt |> break_sep_conj 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2288
                                |> filter (#is_pc_term TP) |> hd
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2289
               val (_$Free(j', _)) = pc_term
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2290
               val psd = post_cond_frame j'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2291
               val str_psd = psd |> string_for_cterm ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2292
               val _ = tracing ctxt "goto_tac: psd = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2293
               val _ = str_psd |> tracing ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2294
           in 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2295
               res_inst_tac ctxt [(("q", 0), str_psd)] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2296
                                          (#post_weaken TP) i state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2297
           end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2298
     val frame_tac = if need_frame then (rtac (#frame_rule TP)) else (K all_tac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2299
     val _ = tracing ctxt "goto_tac: starting to apply tacs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2300
     val print_tac = if (Config.get ctxt trace_step) then Tactical.print_tac else (K all_tac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2301
     val tac = (tac1 THEN' (K (print_tac "tac1 success"))) THEN' 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2302
               (tac2 THEN' (K (print_tac "tac2 success"))) THEN' 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2303
               ((foc_tac post_cond_frame ctxt) THEN' (K (print_tac "foc_tac success"))) THEN' 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2304
               (frame_tac THEN' (K (print_tac "frame_tac success"))) THEN' 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2305
               ((((rtac thm) THEN_ALL_NEW (#sep_conj_ac_tac TP ctxt))) THEN'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2306
                 (K (print_tac "rtac success"))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2307
               ) THEN' 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2308
               (K (ALLGOALS (atomic (#sep_conj_ac_tac TP ctxt)))) THEN'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2309
               (K prune_params_tac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2310
   in 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2311
        tac i state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2312
   end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2313
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2314
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2315
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2316
  fun sg_goto_tac thms ctxt =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2317
     let val sg_goto_tac' =  (map (fn thm  => attemp (goto_tac ctxt thm)) thms)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2318
                              |> FIRST'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2319
         val sg_goto_tac'' = (K (unfold_cell_tac ctxt)) THEN' sg_goto_tac' THEN' (K (fold_cell_tac ctxt))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2320
     in
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2321
         sg_goto_tac' ORELSE' sg_goto_tac''
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2322
     end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2323
  fun gotos_tac thms ctxt i = REPEAT (sg_goto_tac thms ctxt i) THEN (prune_params_tac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2324
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2325
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2326
method_setup hgoto = {* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2327
  Attrib.thms >> (fn thms => fn ctxt =>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2328
                    (SIMPLE_METHOD' (fn i => 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2329
                       sg_goto_tac (thms@(StepRules.get ctxt)) ctxt i)))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2330
  *} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2331
  "One step symbolic execution using goto theorems."
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2332
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2333
subsection {* Tactic for forward reasoning *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2334
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2335
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2336
fun mk_msel_rule ctxt conclusion idx term =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2337
let 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2338
  val cjt_count = term |> break_sep_conj |> length
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2339
  fun variants nctxt names = fold_map Name.variant names nctxt;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2340
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2341
  val (state, nctxt0) = Name.variant "s" (Variable.names_of ctxt);
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2342
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2343
  fun sep_conj_prop cjts =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2344
        FunApp.fun_app_free
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2345
          (FunApp.fun_app_foldr SepConj.sep_conj_term cjts) state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2346
        |> HOLogic.mk_Trueprop;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2347
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2348
  (* concatenate string and string of an int *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2349
  fun conc_str_int str int = str ^ Int.toString int;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2350
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2351
  (* make the conjunct names *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2352
  val (cjts, _) = ListExtra.range 1 cjt_count
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2353
                  |> map (conc_str_int "a") |> variants nctxt0;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2354
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2355
 fun skel_sep_conj names (Const (@{const_name sep_conj}, _) $ t1 $ t2 $ y) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2356
     (let val nm1 = take (length (break_sep_conj t1)) names 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2357
          val nm2 = drop (length (break_sep_conj t1)) names
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2358
          val t1' = skel_sep_conj nm1 t1 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2359
          val t2' = skel_sep_conj nm2 t2 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2360
      in (SepConj.sep_conj_term $ t1' $ t2' $ y) end)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2361
  | skel_sep_conj names (Const (@{const_name sep_conj}, _) $ t1 $ t2) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2362
     (let val nm1 = take (length (break_sep_conj t1)) names 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2363
          val nm2 = drop (length (break_sep_conj t1)) names
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2364
          val t1' = skel_sep_conj nm1 t1 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2365
          val t2' = skel_sep_conj nm2 t2 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2366
     in (SepConj.sep_conj_term $ t1' $ t2') end)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2367
   | skel_sep_conj names (Abs (x, y, t $ Bound 0)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2368
                  let val t' = (skel_sep_conj names t) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2369
                      val ty' = t' |> type_of |> domain_type
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2370
                  in (Abs (x, ty', (t' $ Bound 0))) end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2371
  | skel_sep_conj names t = Free (hd names, SepConj.sep_conj_term |> type_of |> domain_type);
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2372
  val _ = tracing ctxt "xxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxx"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2373
  val oskel = skel_sep_conj cjts term;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2374
  val _ = tracing ctxt "yyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyy"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2375
  val ttt = oskel |> type_of
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2376
  val _ = tracing ctxt "zzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzz"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2377
  val orig = FunApp.fun_app_free oskel state |> HOLogic.mk_Trueprop
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2378
  val _ = tracing ctxt "uuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuu"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2379
  val is_selected = member (fn (x, y) => x = y) idx
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2380
  val all_idx = ListExtra.range 0 cjt_count
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2381
  val selected_idx = idx
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2382
  val unselected_idx = filter_out is_selected all_idx
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2383
  val selected = map (nth cjts) selected_idx
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2384
  val unselected = map (nth cjts) unselected_idx
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2385
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2386
  fun fun_app_foldr f [a,b] = FunApp.fun_app_free (FunApp.fun_app_free f a) b
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2387
  | fun_app_foldr f [a] = Free (a, SepConj.sep_conj_term |> type_of |> domain_type)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2388
  | fun_app_foldr f (x::xs) = (FunApp.fun_app_free f x) $ (fun_app_foldr f xs)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2389
  | fun_app_foldr _ _ = raise Fail "fun_app_foldr";
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2390
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2391
  val reordered_skel = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2392
      if unselected = [] then (fun_app_foldr SepConj.sep_conj_term selected)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2393
          else (SepConj.sep_conj_term $ (fun_app_foldr SepConj.sep_conj_term selected)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2394
                        $ (fun_app_foldr SepConj.sep_conj_term unselected))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2395
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2396
  val reordered =  FunApp.fun_app_free reordered_skel state  |> HOLogic.mk_Trueprop
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2397
  val goal = Logic.mk_implies
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2398
               (if conclusion then (orig, reordered) else (reordered, orig));
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2399
  val rule =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2400
   Goal.prove ctxt [] [] goal (fn _ => 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2401
        auto_tac (ctxt |> map_simpset (fn ss => ss addsimps @{thms sep_conj_ac})))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2402
         |> Drule.export_without_context
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2403
in
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2404
   rule
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2405
end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2406
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2407
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2408
lemma fwd_rule: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2409
  assumes "\<And> s . U s \<longrightarrow> V s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2410
  shows "(U ** RR) s \<Longrightarrow> (V ** RR) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2411
  by (metis assms sep_globalise)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2412
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2413
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2414
  fun sg_sg_fwd_tac ctxt thm pos i state = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2415
  let  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2416
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2417
  val tracing  = (fn str =>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2418
                   if (Config.get ctxt trace_fwd) then Output.tracing str else ())
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2419
  fun pterm t =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2420
          t |> string_of_term ctxt |> tracing
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2421
  fun pcterm ct = ct |> string_of_cterm ctxt |> tracing
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2422
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2423
  fun atm thm = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2424
  let
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2425
  (* val thm = thm |> Drule.forall_intr_vars *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2426
  val res =  thm |> cprop_of |> Object_Logic.atomize
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2427
  val res' = Raw_Simplifier.rewrite_rule [res] thm
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2428
  in res' end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2429
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2430
  fun find_idx ctxt pats terms = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2431
     let val result = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2432
              map (fn pat => (find_index (fn trm => ((match ctxt pat trm; true)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2433
                                              handle _ => false)) terms)) pats
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2434
     in (assert_all (fn x => x >= 0) result (K "match of precondition failed"));
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2435
         result
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2436
     end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2437
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2438
  val goal = nth (Drule.cprems_of state) (i - 1) |> term_of
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2439
  val _ = tracing "goal = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2440
  val _ = goal |> pterm
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2441
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2442
  val ctxt_orig = ctxt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2443
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2444
  val ((ps, goal), ctxt) = Variable.focus goal ctxt_orig
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2445
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2446
  val prems = goal |> Logic.strip_imp_prems 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2447
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2448
  val cprem = nth prems (pos - 1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2449
  val (_ $ (the_prem $ _)) = cprem
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2450
  val cjts = the_prem |> break_sep_conj
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2451
  val thm_prems = thm |> cprems_of |> hd |> Thm.dest_arg |> Thm.dest_fun
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2452
  val thm_assms = thm |> cprems_of |> tl |> map term_of
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2453
  val thm_cjts = thm_prems |> term_of |> break_sep_conj
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2454
  val thm_trm = thm |> prop_of
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2455
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2456
  val _ = tracing "cjts = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2457
  val _ = cjts |> map pterm
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2458
  val _ = tracing "thm_cjts = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2459
  val _ = thm_cjts |> map pterm
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2460
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2461
  val _ = tracing "Calculating sols"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2462
  val sols = match_pres ctxt (match_env ctxt) empty_env thm_cjts cjts
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2463
  val _ = tracing "End calculating sols, sols ="
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2464
  val _ = tracing (@{make_string} sols)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2465
  val _ = tracing "Calulating env2 and idxs"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2466
  val (env2, idx) = filter (fn (env, idxs) => make_sense sep_conj_ac_tac ctxt thm_assms env) sols |> hd
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2467
  val ([thm'_trm], ctxt') = thm_trm |> inst env2 |> single 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2468
                            |> (fn trms => Variable.import_terms true trms ctxt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2469
  val thm'_prem  = Logic.strip_imp_prems thm'_trm |> hd 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2470
  val thm'_concl = Logic.strip_imp_concl thm'_trm 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2471
  val thm'_prem = (Goal.prove ctxt' [] [thm'_prem] thm'_concl 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2472
                  (fn {context, prems = [prem]} =>  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2473
                      (rtac (prem RS thm)  THEN_ALL_NEW (sep_conj_ac_tac ctxt)) 1))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2474
  val [thm'] = Variable.export ctxt' ctxt_orig [thm'_prem]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2475
  val trans_rule = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2476
       mk_msel_rule ctxt true idx the_prem
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2477
  val _ = tracing "trans_rule = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2478
  val _ = trans_rule |> cprop_of |> pcterm
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2479
  val app_rule = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2480
      if (length cjts = length thm_cjts) then thm' else
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2481
       ((thm' |> atm) RS @{thm fwd_rule})
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2482
  val _ = tracing "app_rule = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2483
  val _ = app_rule |> cprop_of |> pcterm
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2484
  val print_tac = if (Config.get ctxt trace_fwd) then Tactical.print_tac else (K all_tac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2485
  val the_tac = (dtac trans_rule THEN' (K (print_tac "dtac1 success"))) THEN'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2486
                ((dtac app_rule THEN' (K (print_tac "dtac2 success"))))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2487
in
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2488
  (the_tac i state) handle _ => no_tac state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2489
end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2490
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2491
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2492
ML {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2493
  fun sg_fwd_tac ctxt thm i state = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2494
  let  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2495
    val goal = nth (Drule.cprems_of state) (i - 1)          
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2496
    val prems = goal |> term_of |> Term.strip_all_body |> Logic.strip_imp_prems 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2497
    val posx = ListExtra.range 1 (length prems)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2498
  in
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2499
      ((map (fn pos => attemp (sg_sg_fwd_tac ctxt thm pos)) posx) |> FIRST') i state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2500
  end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2501
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2502
  fun fwd_tac ctxt thms i state =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2503
       ((map (fn thm => sg_fwd_tac ctxt thm) thms) |> FIRST') i state
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2504
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2505
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2506
method_setup fwd = {* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2507
  Attrib.thms >> (fn thms => fn ctxt =>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2508
                    (SIMPLE_METHOD' (fn i => 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2509
                       fwd_tac ctxt (thms@(FwdRules.get ctxt))  i)))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2510
  *} 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2511
  "Forward derivation of separation implication"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2512
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2513
text {* Testing the fwd tactic *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2514
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2515
lemma ones_abs:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2516
  assumes "(ones u v \<and>* ones w x) s" "w = v + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2517
  shows "ones u x s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2518
  using assms(1) unfolding assms(2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2519
proof(induct u v arbitrary: x s rule:ones_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2520
  case (Base i j x s)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2521
  thus ?case by (auto elim!:condE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2522
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2523
  case (Step i j x s)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2524
  hence h: "\<And> x s. (ones (i + 1) j \<and>* ones (j + 1) x) s \<longrightarrow> ones (i + 1) x s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2525
    by metis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2526
  hence "(ones (i + 1) x \<and>* one i) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2527
    by (rule fwd_rule, insert Step(3), auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2528
  thus ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2529
    by (smt condD ones.simps sep_conj_commute)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2530
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2531
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2532
lemma one_abs: "(one m) s \<Longrightarrow> (ones m m) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2533
 by (smt cond_true_eq2 ones.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2534
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2535
lemma ones_reps_abs: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2536
  assumes "ones m n s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2537
          "m \<le> n"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2538
  shows "(reps m n [nat (n - m)]) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2539
  using assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2540
  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2541
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2542
lemma reps_reps'_abs: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2543
  assumes "(reps m n xs \<and>* zero u) s" "u = n + 1" "xs \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2544
  shows "(reps' m u xs) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2545
  unfolding assms using assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2546
  by (unfold reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2547
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2548
lemma reps'_abs:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2549
  assumes "(reps' m n xs \<and>* reps' u v ys) s" "u = n + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2550
  shows "(reps' m v (xs @ ys)) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2551
  apply (unfold reps'_append, rule_tac x = u in EXS_intro)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2552
  by (insert assms, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2553
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2554
lemmas abs_ones = one_abs ones_abs
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2555
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2556
lemmas abs_reps' = ones_reps_abs reps_reps'_abs reps'_abs
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2557
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2558
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2559
section {* Modular TM programming and verification *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2560
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2561
definition "right_until_zero = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2562
                 (TL start exit. 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2563
                  TLabel start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2564
                     if_zero exit;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2565
                     move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2566
                     jmp start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2567
                  TLabel exit
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2568
                 )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2569
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2570
lemma ones_false [simp]: "j < i - 1 \<Longrightarrow> (ones i j) = sep_false"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2571
  by (simp add:pasrt_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2572
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2573
lemma hoare_right_until_zero: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2574
  "\<lbrace>st i ** ps u ** ones u (v - 1) ** zero v \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2575
     i:[right_until_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2576
   \<lbrace>st j ** ps v ** ones u (v - 1) ** zero v \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2577
proof(unfold right_until_zero_def, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2578
      intro t_hoare_local t_hoare_label, clarify, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2579
      rule t_hoare_label_last, simp, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2580
  fix la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2581
  let ?body = "i :[ (if_zero la ; move_right ; jmp i) ]: la"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2582
  let ?j = la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2583
  show "\<lbrace>st i \<and>* ps u \<and>* ones u (v - 1) \<and>* zero v\<rbrace>  ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2584
        \<lbrace>st ?j \<and>* ps v \<and>* ones u (v - 1) \<and>* zero v\<rbrace>" (is "?P u (v - 1) (ones u (v - 1))")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2585
  proof(induct "u" "v - 1" rule:ones_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2586
    case (Base k)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2587
    moreover have "\<lbrace>st i \<and>* ps v \<and>* zero v\<rbrace> ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2588
                   \<lbrace>st ?j \<and>* ps v \<and>* zero v\<rbrace>" by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2589
    ultimately show ?case by (auto intro!:tm.pre_condI simp:sep_conj_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2590
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2591
    case (Step k)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2592
    moreover have "\<lbrace>st i \<and>* ps k \<and>* (one k \<and>* ones (k + 1) (v - 1)) \<and>* zero v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2593
                     i :[ (if_zero ?j ; move_right ; jmp i) ]: ?j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2594
                   \<lbrace>st ?j \<and>* ps v \<and>* (one k \<and>* ones (k + 1) (v - 1)) \<and>* zero v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2595
    proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2596
      have s1: "\<lbrace>st i \<and>* ps k \<and>* (one k \<and>* ones (k + 1) (v - 1)) \<and>* zero v\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2597
                          ?body 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2598
                \<lbrace>st i \<and>* ps (k + 1) \<and>* one k \<and>* ones (k + 1) (v - 1) \<and>* zero v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2599
      proof(cases "k + 1 \<ge> v")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2600
        case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2601
        with Step(1) have "v = k + 1" by arith
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2602
        thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2603
          apply(simp add: one_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2604
          by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2605
      next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2606
        case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2607
        hence eq_ones: "ones (k + 1) (v - 1) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2608
                         (one (k + 1) \<and>* ones ((k + 1) + 1) (v - 1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2609
          by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2610
        show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2611
          apply(simp only: eq_ones)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2612
          by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2613
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2614
      note Step(2)[step]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2615
      have s2: "\<lbrace>st i \<and>* ps (k + 1) \<and>* one k \<and>* ones (k + 1) (v - 1) \<and>* zero v\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2616
                        ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2617
                \<lbrace>st ?j \<and>* ps v \<and>* one k \<and>* ones (k + 1) (v - 1) \<and>* zero v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2618
        by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2619
      from tm.sequencing [OF s1 s2, step] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2620
      show ?thesis 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2621
        by (auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2622
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2623
    ultimately show ?case by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2624
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2625
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2626
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2627
lemma hoare_right_until_zero_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2628
  assumes "u = v" "w = x - 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2629
  shows  "\<lbrace>st i ** ps u ** ones v w ** zero x \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2630
              i:[right_until_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2631
          \<lbrace>st j ** ps x ** ones v w ** zero x \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2632
  by (unfold assms, rule hoare_right_until_zero)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2633
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2634
definition "left_until_zero = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2635
                 (TL start exit. 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2636
                  TLabel start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2637
                    if_zero exit;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2638
                    move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2639
                    jmp start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2640
                  TLabel exit
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2641
                 )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2642
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2643
lemma hoare_left_until_zero: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2644
  "\<lbrace>st i ** ps v ** zero u ** ones (u + 1) v \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2645
     i:[left_until_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2646
   \<lbrace>st j ** ps u ** zero u ** ones (u + 1) v \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2647
proof(unfold left_until_zero_def, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2648
      intro t_hoare_local t_hoare_label, clarify, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2649
      rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2650
  fix la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2651
  let ?body = "i :[ (if_zero la ; move_left ; jmp i) ]: la"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2652
  let ?j = la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2653
  show "\<lbrace>st i \<and>* ps v \<and>* zero u \<and>* ones (u + 1) v\<rbrace> ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2654
        \<lbrace>st ?j \<and>* ps u \<and>* zero u \<and>* ones (u + 1) v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2655
  proof(induct "u+1" v  rule:ones_rev_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2656
    case (Base k)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2657
    thus ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2658
      by (simp add:sep_conj_cond, intro tm.pre_condI, simp, hstep)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2659
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2660
    case (Step k)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2661
    have "\<lbrace>st i \<and>* ps k \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2662
               ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2663
          \<lbrace>st ?j \<and>* ps u \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2664
    proof(rule tm.sequencing[where q = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2665
           "st i \<and>* ps (k - 1) \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2666
      show "\<lbrace>st i \<and>* ps k \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2667
                ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2668
            \<lbrace>st i \<and>* ps (k - 1) \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2669
      proof(induct "u + 1" "k - 1" rule:ones_rev_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2670
        case Base with Step(1) have "k = u + 1" by arith
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2671
        thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2672
          by (simp, hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2673
      next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2674
        case Step
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2675
        show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2676
          apply (unfold ones_rev[OF Step(1)], simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2677
          apply (unfold one_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2678
          by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2679
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2680
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2681
      note Step(2) [step]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2682
      show "\<lbrace>st i \<and>* ps (k - 1) \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2683
                ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2684
            \<lbrace>st ?j \<and>* ps u \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace>" by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2685
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2686
    thus ?case by (unfold ones_rev[OF Step(1)], simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2687
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2688
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2689
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2690
lemma hoare_left_until_zero_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2691
  assumes "u = x" "w = v + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2692
  shows  "\<lbrace>st i ** ps u ** zero v ** ones w x \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2693
               i:[left_until_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2694
          \<lbrace>st j ** ps v ** zero v ** ones w x \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2695
  by (unfold assms, rule hoare_left_until_zero)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2696
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2697
definition "right_until_one = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2698
                 (TL start exit. 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2699
                  TLabel start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2700
                     if_one exit;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2701
                     move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2702
                     jmp start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2703
                  TLabel exit
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2704
                 )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2705
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2706
lemma hoare_right_until_one: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2707
  "\<lbrace>st i ** ps u ** zeros u (v - 1) ** one v \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2708
     i:[right_until_one]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2709
   \<lbrace>st j ** ps v ** zeros u (v - 1) ** one v \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2710
proof(unfold right_until_one_def, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2711
      intro t_hoare_local t_hoare_label, clarify, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2712
      rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2713
  fix la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2714
  let ?body = "i :[ (if_one la ; move_right ; jmp i) ]: la"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2715
  let ?j = la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2716
  show "\<lbrace>st i \<and>* ps u \<and>* zeros u (v - 1) \<and>* one v\<rbrace> ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2717
       \<lbrace>st ?j \<and>* ps v \<and>* zeros u (v - 1) \<and>* one v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2718
  proof(induct u "v - 1" rule:zeros_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2719
    case (Base k)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2720
    thus ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2721
      by (simp add:sep_conj_cond, intro tm.pre_condI, simp, hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2722
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2723
    case (Step k)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2724
    have "\<lbrace>st i \<and>* ps k \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2725
            ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2726
          \<lbrace>st ?j \<and>* ps v \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2727
    proof(rule tm.sequencing[where q = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2728
           "st i \<and>* ps (k + 1) \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2729
      show "\<lbrace>st i \<and>* ps k \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2730
               ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2731
           \<lbrace>st i \<and>* ps (k + 1) \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2732
      proof(induct "k + 1" "v - 1" rule:zeros_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2733
        case Base
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2734
        with Step(1) have eq_v: "k + 1 = v" by arith
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2735
        from Base show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2736
          apply (simp add:sep_conj_cond, intro tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2737
          apply (hstep, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2738
          by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2739
      next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2740
        case Step
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2741
        thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2742
          by (simp, hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2743
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2744
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2745
      note Step(2)[step]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2746
        show "\<lbrace>st i \<and>* ps (k + 1) \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2747
                ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2748
              \<lbrace>st ?j \<and>* ps v \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2749
          by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2750
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2751
    thus ?case by (auto simp: sep_conj_ac Step(1))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2752
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2753
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2754
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2755
lemma hoare_right_until_one_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2756
  assumes "u = v" "w = x - 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2757
  shows
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2758
  "\<lbrace>st i ** ps u ** zeros v w ** one x \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2759
     i:[right_until_one]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2760
   \<lbrace>st j **  ps x ** zeros v w ** one x \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2761
  by (unfold assms, rule hoare_right_until_one)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2762
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2763
definition "left_until_one = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2764
                 (TL start exit. 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2765
                  TLabel start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2766
                    if_one exit;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2767
                    move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2768
                    jmp start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2769
                  TLabel exit
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2770
                 )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2771
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2772
lemma hoare_left_until_one: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2773
  "\<lbrace>st i ** ps v ** one u ** zeros (u + 1) v \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2774
     i:[left_until_one]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2775
   \<lbrace>st j ** ps u ** one u ** zeros (u + 1) v \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2776
proof(unfold left_until_one_def, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2777
      intro t_hoare_local t_hoare_label, clarify, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2778
      rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2779
  fix la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2780
  let ?body = "i :[ (if_one la ; move_left ; jmp i) ]: la"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2781
  let ?j = la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2782
  show "\<lbrace>st i \<and>* ps v \<and>* one u \<and>* zeros (u + 1) v\<rbrace> ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2783
        \<lbrace>st ?j \<and>* ps u \<and>* one u \<and>* zeros (u + 1) v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2784
  proof(induct u v rule: ones'.induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2785
    fix ia ja
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2786
    assume h: "\<not> ja < ia \<Longrightarrow>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2787
             \<lbrace>st i \<and>* ps (ja - 1) \<and>* one ia \<and>* zeros (ia + 1) (ja - 1)\<rbrace> ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2788
             \<lbrace>st ?j \<and>* ps ia \<and>* one ia \<and>* zeros (ia + 1) (ja - 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2789
    show "\<lbrace>st i \<and>* ps ja \<and>* one ia \<and>* zeros (ia + 1) ja\<rbrace>  ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2790
      \<lbrace>st ?j \<and>* ps ia \<and>* one ia \<and>* zeros (ia + 1) ja\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2791
    proof(cases "ja < ia")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2792
      case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2793
      note lt = False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2794
      from h[OF this] have [step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2795
        "\<lbrace>st i \<and>* ps (ja - 1) \<and>* one ia \<and>* zeros (ia + 1) (ja - 1)\<rbrace> ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2796
         \<lbrace>st ?j \<and>* ps ia \<and>* one ia \<and>* zeros (ia + 1) (ja - 1)\<rbrace>" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2797
      show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2798
      proof(cases "ja = ia")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2799
        case True 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2800
        moreover
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2801
        have "\<lbrace>st i \<and>* ps ja \<and>* one ja\<rbrace> ?body \<lbrace>st ?j \<and>* ps ja \<and>* one ja\<rbrace>" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2802
          by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2803
        ultimately show ?thesis by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2804
      next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2805
        case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2806
        with lt have k1: "ia < ja" by auto       
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2807
        from zeros_rev[of "ja" "ia + 1"] this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2808
        have eq_zeros: "zeros (ia + 1) ja = (zeros (ia + 1) (ja - 1) \<and>* zero ja)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2809
          by simp        
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2810
        have s1: "\<lbrace>st i \<and>* ps ja \<and>* one ia \<and>* zeros (ia + 1) (ja - 1) \<and>* zero ja\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2811
                      ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2812
                  \<lbrace>st i \<and>* ps (ja - 1) \<and>* one ia \<and>* zeros (ia + 1) (ja - 1) \<and>* zero ja\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2813
        proof(cases "ia + 1 \<ge> ja")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2814
          case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2815
          from k1 True have "ja = ia + 1" by arith
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2816
          moreover have "\<lbrace>st i \<and>* ps (ia + 1) \<and>* one (ia + 1 - 1) \<and>* zero (ia + 1)\<rbrace>  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2817
            i :[ (if_one ?j ; move_left ; jmp i) ]: ?j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2818
                \<lbrace>st i \<and>* ps (ia + 1 - 1) \<and>* one (ia + 1 - 1) \<and>* zero (ia + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2819
            by (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2820
          ultimately show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2821
            by (simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2822
        next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2823
          case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2824
          from zeros_rev[of "ja - 1" "ia + 1"] False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2825
          have k: "zeros (ia + 1) (ja - 1) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2826
                      (zeros (ia + 1) (ja - 1 - 1) \<and>* zero (ja - 1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2827
            by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2828
          show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2829
            apply (unfold k, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2830
            by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2831
        qed      
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2832
        have s2: "\<lbrace>st i \<and>* ps (ja - 1) \<and>* one ia \<and>* zeros (ia + 1) (ja - 1) \<and>* zero ja\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2833
                      ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2834
                  \<lbrace>st ?j \<and>* ps ia \<and>* one ia \<and>* zeros (ia + 1) (ja - 1) \<and>* zero ja\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2835
          by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2836
        from tm.sequencing[OF s1 s2, step]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2837
        show ?thesis 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2838
          apply (unfold eq_zeros)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2839
          by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2840
      qed (* ccc *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2841
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2842
      case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2843
      thus ?thesis by (auto intro:tm.hoare_sep_false)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2844
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2845
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2846
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2847
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2848
lemma hoare_left_until_one_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2849
  assumes "u = x" "w = v + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2850
  shows  "\<lbrace>st i ** ps u ** one v ** zeros w x \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2851
              i:[left_until_one]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2852
          \<lbrace>st j ** ps v ** one v ** zeros w x \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2853
  by (unfold assms, rule hoare_left_until_one)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2854
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2855
definition "left_until_double_zero = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2856
            (TL start exit.
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2857
              TLabel start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2858
              if_zero exit;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2859
              left_until_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2860
              move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2861
              if_one start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2862
              TLabel exit)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2863
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2864
declare ones.simps[simp del]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2865
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2866
lemma reps_simps3: "ks \<noteq> [] \<Longrightarrow> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2867
  reps i j (k # ks) = (ones i (i + int k) ** zero (i + int k + 1) ** reps (i + int k + 2) j ks)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2868
by(case_tac ks, simp, simp add: reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2869
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2870
lemma cond_eqI:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2871
  assumes h: "b \<Longrightarrow> r = s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2872
  shows "(<b> ** r) = (<b> ** s)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2873
proof(cases b)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2874
  case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2875
  from h[OF this] show ?thesis by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2876
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2877
  case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2878
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2879
    by (unfold sep_conj_def set_ins_def pasrt_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2880
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2881
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2882
lemma reps_rev: "ks \<noteq> [] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2883
       \<Longrightarrow> reps i j (ks @ [k]) =  (reps i (j - int (k + 1) - 1 ) ks \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2884
                                          zero (j - int (k + 1)) \<and>* ones (j - int k) j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2885
proof(induct ks arbitrary: i j)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2886
  case Nil
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2887
  thus ?case by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2888
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2889
  case (Cons a ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2890
  show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2891
  proof(cases "ks = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2892
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2893
    thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2894
    proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2895
      have eq_cond: "(j = i + int a + 2 + int k) = (-2 + (j - int k) = i + int a)" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2896
      have "(<(-2 + (j - int k) = i + int a)> \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2897
            one i \<and>* ones (i + 1) (i + int a) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2898
            zero (i + int a + 1) \<and>* one (i + int a + 2) \<and>* ones (3 + (i + int a)) (i + int a + 2 + int k))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2899
            =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2900
            (<(-2 + (j - int k) = i + int a)> \<and>* one i \<and>* ones (i + 1) (i + int a) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2901
            zero (j - (1 + int k)) \<and>* one (j - int k) \<and>* ones (j - int k + 1) j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2902
        (is "(<?X> \<and>* ?L) = (<?X> \<and>* ?R)")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2903
      proof(rule cond_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2904
        assume h: "-2 + (j - int k) = i + int a"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2905
        hence eqs:  "i + int a + 1 = j - (1 + int k)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2906
                    "i + int a + 2 = j - int k"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2907
                    "3 + (i + int a) = j - int k + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2908
                    "(i + int a + 2 + int k) = j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2909
        by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2910
        show "?L = ?R"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2911
          by (unfold eqs, auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2912
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2913
      with True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2914
      show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2915
        apply (simp del:ones_simps reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2916
        apply (simp add:sep_conj_cond eq_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2917
        by (auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2918
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2919
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2920
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2921
    from Cons(1)[OF False, of "i + int a + 2" j] this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2922
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2923
      by(simp add: reps_simps3 sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2924
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2925
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2926
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2927
lemma hoare_if_one_reps:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2928
  assumes nn: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2929
  shows "\<lbrace>st i ** ps v ** reps u v ks\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2930
           i:[if_one e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2931
        \<lbrace>st e ** ps v ** reps u v ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2932
proof(rule rev_exhaust[of ks])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2933
  assume "ks = []" with nn show ?thesis by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2934
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2935
  fix y ys
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2936
  assume eq_ks: "ks = ys @ [y]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2937
  show " \<lbrace>st i \<and>* ps v \<and>* reps u v ks\<rbrace>  i :[ if_one e ]: j \<lbrace>st e \<and>* ps v \<and>* reps u v ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2938
  proof(cases "ys = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2939
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2940
    have "\<lbrace>st i \<and>* ps v \<and>* reps u v (ys @ [y])\<rbrace>  i :[ if_one e ]: j \<lbrace>st e \<and>* ps v \<and>* reps u v (ys @ [y])\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2941
      apply(unfold reps_rev[OF False], simp del:ones_simps add:ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2942
      by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2943
    thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2944
      by (simp add:eq_ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2945
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2946
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2947
    with eq_ks
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2948
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2949
      apply (simp del:ones_simps add:ones_rev sep_conj_cond, intro tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2950
      by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2951
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2952
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2953
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2954
lemma hoare_if_one_reps_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2955
  assumes nn: "ks \<noteq> []" "u = w"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2956
  shows "\<lbrace>st i ** ps u ** reps v w ks\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2957
           i:[if_one e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2958
        \<lbrace>st e ** ps u ** reps v w ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2959
  by (unfold `u = w`, rule hoare_if_one_reps[OF `ks \<noteq> []`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2960
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2961
lemma hoare_if_zero_ones_false[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2962
  assumes "\<not> w < u" "v = w"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2963
  shows  "\<lbrace>st i \<and>* ps v \<and>* ones u w\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2964
             i :[if_zero e]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2965
          \<lbrace>st j \<and>* ps v \<and>* ones u w\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2966
  by (unfold `v = w` ones_rev[OF `\<not> w < u`], hstep)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2967
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2968
lemma hoare_left_until_double_zero_nil[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2969
  assumes "u = v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2970
  shows   "\<lbrace>st i ** ps u ** zero v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2971
                  i:[left_until_double_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2972
           \<lbrace>st j ** ps u ** zero v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2973
  apply (unfold `u = v` left_until_double_zero_def, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2974
      intro t_hoare_local t_hoare_label, clarsimp, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2975
      rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2976
  by (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2977
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2978
lemma hoare_if_zero_reps_false:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2979
  assumes nn: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2980
  shows "\<lbrace>st i ** ps v ** reps u v ks\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2981
           i:[if_zero e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2982
        \<lbrace>st j ** ps v ** reps u v ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2983
proof(rule rev_exhaust[of ks])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2984
  assume "ks = []" with nn show ?thesis by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2985
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2986
  fix y ys
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2987
  assume eq_ks: "ks = ys @ [y]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2988
  show " \<lbrace>st i \<and>* ps v \<and>* reps u v ks\<rbrace>  i :[ if_zero e ]: j \<lbrace>st j \<and>* ps v \<and>* reps u v ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2989
  proof(cases "ys = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2990
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2991
    have "\<lbrace>st i \<and>* ps v \<and>* reps u v (ys @ [y])\<rbrace>  i :[ if_zero e ]: j \<lbrace>st j \<and>* ps v \<and>* reps u v (ys @ [y])\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2992
      apply(unfold reps_rev[OF False], simp del:ones_simps add:ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2993
      by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2994
    thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2995
      by (simp add:eq_ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2996
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2997
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2998
    with eq_ks
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  2999
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3000
      apply (simp del:ones_simps add:ones_rev sep_conj_cond, intro tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3001
      by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3002
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3003
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3004
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3005
lemma hoare_if_zero_reps_false_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3006
  assumes "ks \<noteq> []" "u = w"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3007
  shows "\<lbrace>st i ** ps u ** reps v w ks\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3008
           i:[if_zero e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3009
        \<lbrace>st j ** ps u ** reps v w ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3010
  by (unfold `u = w`, rule hoare_if_zero_reps_false[OF `ks \<noteq> []`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3011
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3012
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3013
lemma hoare_if_zero_reps_false1:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3014
  assumes nn: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3015
  shows "\<lbrace>st i ** ps u ** reps u v ks\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3016
           i:[if_zero e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3017
        \<lbrace>st j ** ps u ** reps u v ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3018
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3019
  from nn obtain y ys where eq_ys: "ks = y#ys"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3020
    by (metis neq_Nil_conv)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3021
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3022
    apply (unfold eq_ys)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3023
    by (case_tac ys, (simp, hsteps)+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3024
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3025
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3026
lemma hoare_if_zero_reps_false1_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3027
  assumes nn: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3028
  and h: "u = w"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3029
  shows "\<lbrace>st i ** ps u ** reps w v ks\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3030
           i:[if_zero e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3031
        \<lbrace>st j ** ps u ** reps w v ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3032
  by (unfold h, rule hoare_if_zero_reps_false1[OF `ks \<noteq> []`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3033
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3034
lemma hoare_left_until_double_zero: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3035
  assumes h: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3036
  shows   "\<lbrace>st i ** ps v ** zero u ** zero (u + 1) ** reps (u+2) v ks\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3037
                  i:[left_until_double_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3038
           \<lbrace>st j ** ps u ** zero u ** zero (u + 1) ** reps (u+2) v ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3039
proof(unfold left_until_double_zero_def, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3040
      intro t_hoare_local t_hoare_label, clarsimp, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3041
      rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3042
  fix la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3043
  let ?body = "i :[ (if_zero la ; left_until_zero ; move_left ; if_one i) ]: j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3044
  let ?j = j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3045
  show "\<lbrace>st i \<and>* ps v \<and>* zero u \<and>* zero (u + 1) \<and>* reps (u + 2) v ks\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3046
           ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3047
        \<lbrace>st ?j \<and>* ps u \<and>* zero u \<and>* zero (u + 1) \<and>* reps (u + 2) v ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3048
    using h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3049
  proof(induct ks arbitrary: v rule:rev_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3050
    case Nil
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3051
    with h show ?case by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3052
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3053
    case (snoc k ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3054
    show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3055
    proof(cases "ks = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3056
      case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3057
      have eq_ones: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3058
        "ones (u + 2) (u + 2 + int k) = (ones (u + 2) (u + 1 + int k) \<and>* one (u + 2 + int k))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3059
        by (smt ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3060
      have eq_ones': "(one (u + 2) \<and>* ones (3 + u) (u + 2 + int k)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3061
            (one (u + 2 + int k) \<and>* ones (u + 2) (u + 1 + int k))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3062
        by (smt eq_ones ones.simps sep.mult_commute)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3063
      thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3064
        apply (insert True, simp del:ones_simps add:sep_conj_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3065
        apply (rule tm.pre_condI, simp del:ones_simps, unfold eq_ones)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3066
        apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3067
        apply (rule_tac p = "st j' \<and>* ps (u + 2 + int k) \<and>* zero u \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3068
                             zero (u + 1) \<and>* ones (u + 2) (u + 2 + int k)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3069
                  in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3070
        by (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3071
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3072
      case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3073
      from False have spt: "splited (ks @ [k]) ks [k]" by (unfold splited_def, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3074
      show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3075
        apply (unfold reps_splited[OF spt], simp del:ones_simps add:sep_conj_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3076
        apply (rule tm.pre_condI, simp del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3077
        apply (rule_tac q = "st i \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3078
               ps (1 + (u + int (reps_len ks))) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3079
               zero u \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3080
               zero (u + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3081
               reps (u + 2) (1 + (u + int (reps_len ks))) ks \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3082
               zero (u + 2 + int (reps_len ks)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3083
               ones (3 + (u + int (reps_len ks))) (3 + (u + int (reps_len ks)) + int k)" in
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3084
               tm.sequencing)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3085
        apply hsteps[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3086
        by (hstep snoc(1))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3087
    qed 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3088
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3089
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3090
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3091
lemma hoare_left_until_double_zero_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3092
  assumes h1: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3093
      and h: "u = y" "w = v + 1" "x = v + 2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3094
  shows   "\<lbrace>st i ** ps u ** zero v ** zero w ** reps x y ks\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3095
                  i:[left_until_double_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3096
           \<lbrace>st j ** ps v ** zero v ** zero w ** reps x y ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3097
  by (unfold h, rule hoare_left_until_double_zero[OF h1])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3098
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3099
lemma hoare_jmp_reps1:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3100
  assumes "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3101
  shows  "\<lbrace> st i \<and>* ps u \<and>* reps u v ks\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3102
                 i:[jmp e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3103
          \<lbrace> st e \<and>* ps u \<and>* reps u v ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3104
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3105
  from assms obtain k ks' where Cons:"ks = k#ks'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3106
    by (metis neq_Nil_conv)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3107
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3108
  proof(cases "ks' = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3109
    case True with Cons
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3110
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3111
      apply(simp add:sep_conj_cond reps.simps, intro tm.pre_condI, simp add:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3112
      by (hgoto hoare_jmp_gen)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3113
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3114
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3115
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3116
      apply (unfold `ks = k#ks'` reps_simp3[OF False], simp add:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3117
      by (hgoto hoare_jmp[where p = u])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3118
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3119
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3120
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3121
lemma hoare_jmp_reps1_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3122
  assumes "ks \<noteq> []" "u = v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3123
  shows  "\<lbrace> st i \<and>* ps u \<and>* reps v w ks\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3124
                 i:[jmp e]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3125
          \<lbrace> st e \<and>* ps u \<and>* reps v w ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3126
  by (unfold assms, rule hoare_jmp_reps1[OF `ks \<noteq> []`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3127
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3128
lemma hoare_jmp_reps:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3129
      "\<lbrace> st i \<and>* ps u \<and>* reps u v ks \<and>* tm (v + 1) x \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3130
                 i:[(jmp e; c)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3131
       \<lbrace> st e \<and>* ps u \<and>* reps u v ks \<and>* tm (v + 1) x \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3132
proof(cases "ks")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3133
  case Nil
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3134
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3135
    by (simp add:sep_conj_cond, intro tm.pre_condI, simp, hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3136
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3137
  case (Cons k ks')
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3138
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3139
  proof(cases "ks' = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3140
    case True with Cons
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3141
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3142
      apply(simp add:sep_conj_cond, intro tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3143
      by (hgoto hoare_jmp[where p = u])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3144
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3145
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3146
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3147
      apply (unfold `ks = k#ks'` reps_simp3[OF False], simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3148
      by (hgoto hoare_jmp[where p = u])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3149
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3150
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3151
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3152
definition "shift_right =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3153
            (TL start exit.
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3154
              TLabel start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3155
                 if_zero exit;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3156
                 write_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3157
                 move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3158
                 right_until_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3159
                 write_one;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3160
                 move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3161
                 jmp start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3162
              TLabel exit
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3163
            )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3164
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3165
lemma hoare_shift_right_cons:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3166
  assumes h: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3167
  shows "\<lbrace>st i \<and>* ps u ** reps u v ks \<and>* zero (v + 1) \<and>* zero (v + 2) \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3168
            i:[shift_right]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3169
         \<lbrace>st j ** ps (v + 2) ** zero u ** reps (u + 1) (v + 1) ks ** zero (v + 2) \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3170
proof(unfold shift_right_def, intro t_hoare_local t_hoare_label, clarify, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3171
      rule t_hoare_label_last, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3172
  fix la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3173
  have eq_ones: "\<And> u k. (one (u + int k + 1) \<and>* ones (u + 1) (u + int k)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3174
                                   (one (u + 1) \<and>* ones (2 + u) (u + 1 + int k))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3175
    by (smt cond_true_eq2 ones.simps ones_rev sep.mult_assoc sep.mult_commute 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3176
               sep.mult_left_commute sep_conj_assoc sep_conj_commute 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3177
               sep_conj_cond1 sep_conj_cond2 sep_conj_cond3 sep_conj_left_commute
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3178
               sep_conj_trivial_strip2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3179
  show "\<lbrace>st i \<and>* ps u \<and>* reps u v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3180
         i :[ (if_zero la ;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3181
               write_zero ; move_right ; right_until_zero ; write_one ; move_right ; jmp i) ]: la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3182
         \<lbrace>st la \<and>* ps (v + 2) \<and>* zero u \<and>* reps (u + 1) (v + 1) ks \<and>* zero (v + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3183
    using h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3184
  proof(induct ks arbitrary:i u v)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3185
    case (Cons k ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3186
    thus ?case 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3187
    proof(cases "ks = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3188
      let ?j = la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3189
      case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3190
      let ?body = "i :[ (if_zero ?j ;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3191
                      write_zero ;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3192
                      move_right ; 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3193
                      right_until_zero ; 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3194
                      write_one ; move_right ; jmp i) ]: ?j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3195
      have first_interation: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3196
           "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3197
                                                                             zero (u + int k + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3198
                ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3199
            \<lbrace>st i \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3200
             ps (u + int k + 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3201
             one (u + int k + 1) \<and>* ones (u + 1) (u + int k) \<and>* zero u \<and>* zero (u + int k + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3202
        apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3203
        by (simp add:sep_conj_ac, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3204
      hence "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3205
                                                                             zero (u + int k + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3206
                   ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3207
             \<lbrace>st ?j \<and>* ps (u + int k + 2) \<and>* zero u \<and>* one (u + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3208
                         ones (2 + u) (u + 1 + int k) \<and>* zero (u + int k + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3209
      proof(rule tm.sequencing)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3210
        show "\<lbrace>st i \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3211
               ps (u + int k + 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3212
               one (u + int k + 1) \<and>* ones (u + 1) (u + int k) \<and>* zero u \<and>* zero (u + int k + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3213
                      ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3214
              \<lbrace>st ?j \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3215
               ps (u + int k + 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3216
               zero u \<and>* one (u + 1) \<and>* ones (2 + u) (u + 1 + int k) \<and>* zero (u + int k + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3217
          apply (hgoto hoare_if_zero_true_gen)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3218
          by (simp add:sep_conj_ac eq_ones)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3219
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3220
      with True 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3221
      show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3222
        by (simp, simp only:sep_conj_cond, intro tm.pre_condI, auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3223
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3224
      case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3225
      let ?j = la
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3226
      let ?body = "i :[ (if_zero ?j ;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3227
                        write_zero ;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3228
                        move_right ; right_until_zero ; 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3229
                        write_one ; move_right ; jmp i) ]: ?j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3230
      have eq_ones': 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3231
         "(one (u + int k + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3232
           ones (u + 1) (u + int k) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3233
           zero u \<and>* reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3234
                   =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3235
           (zero u \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3236
             ones (u + 1) (u + int k) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3237
             one (u + int k + 1) \<and>* reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3238
        by (simp add:eq_ones sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3239
      have "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3240
                 reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3241
                    ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3242
            \<lbrace>st i \<and>* ps (u + int k + 2) \<and>* zero u \<and>* ones (u + 1) (u + int k) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3243
                 one (u + int k + 1) \<and>* reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3244
        apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3245
        by (auto simp:sep_conj_ac, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3246
      hence "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3247
                 reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3248
                     ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3249
            \<lbrace>st ?j \<and>* ps (v + 2) \<and>* zero u \<and>* one (u + 1) \<and>* ones (2 + u) (u + 1 + int k) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3250
                 zero (2 + (u + int k)) \<and>* reps (3 + (u + int k)) (v + 1) ks \<and>* zero (v + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3251
      proof(rule tm.sequencing)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3252
        have eq_ones': 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3253
          "\<And> u k. (one (u + int k + 1) \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 2)) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3254
             (one (u + 1) \<and>* zero (2 + (u + int k)) \<and>* ones (2 + u) (u + 1 + int k))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3255
          by (smt eq_ones sep.mult_assoc sep_conj_commute)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3256
        show "\<lbrace>st i \<and>* ps (u + int k + 2) \<and>* zero u \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3257
                    ones (u + 1) (u + int k) \<and>* one (u + int k + 1) \<and>* reps (u + int k + 2) v ks \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3258
                    zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3259
                      ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3260
              \<lbrace>st ?j \<and>* ps (v + 2) \<and>* zero u \<and>* one (u + 1) \<and>* ones (2 + u) (u + 1 + int k) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3261
                      zero (2 + (u + int k)) \<and>* reps (3 + (u + int k)) (v + 1) ks \<and>* zero (v + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3262
          apply (hsteps Cons.hyps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3263
          by (simp add:sep_conj_ac eq_ones, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3264
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3265
      thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3266
        by (unfold reps_simp3[OF False], auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3267
    qed 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3268
  qed auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3269
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3270
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3271
lemma hoare_shift_right_cons_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3272
  assumes h: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3273
  and h1: "u = v" "x = w + 1" "y = w + 2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3274
  shows "\<lbrace>st i \<and>* ps u ** reps v w ks \<and>* zero x \<and>* zero y \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3275
            i:[shift_right]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3276
         \<lbrace>st j ** ps y ** zero v ** reps (v + 1) x ks ** zero y\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3277
  by (unfold h1, rule hoare_shift_right_cons[OF h])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3278
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3279
lemma shift_right_nil [step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3280
  assumes "u = v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3281
  shows
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3282
       "\<lbrace> st i \<and>* ps u \<and>* zero v \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3283
               i:[shift_right]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3284
        \<lbrace> st j \<and>* ps u \<and>* zero v \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3285
  by (unfold assms shift_right_def, intro t_hoare_local t_hoare_label, clarify, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3286
          rule t_hoare_label_last, simp+, hstep)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3287
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3288
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3289
text {*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3290
  @{text "clear_until_zero"} is useful to implement @{text "drag"}.
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3291
*}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3292
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3293
definition "clear_until_zero = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3294
             (TL start exit.
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3295
              TLabel start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3296
                 if_zero exit;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3297
                 write_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3298
                 move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3299
                 jmp start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3300
              TLabel exit)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3301
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3302
lemma  hoare_clear_until_zero[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3303
         "\<lbrace>st i ** ps u ** ones u v ** zero (v + 1)\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3304
              i :[clear_until_zero]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3305
          \<lbrace>st j ** ps (v + 1) ** zeros u v ** zero (v + 1)\<rbrace> "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3306
proof(unfold clear_until_zero_def, intro t_hoare_local, rule t_hoare_label,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3307
    rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3308
  let ?body = "i :[ (if_zero j ; write_zero ; move_right ; jmp i) ]: j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3309
  show "\<lbrace>st i \<and>* ps u \<and>* ones u v \<and>* zero (v + 1)\<rbrace> ?body 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3310
        \<lbrace>st j \<and>* ps (v + 1) \<and>* zeros u v \<and>* zero (v + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3311
  proof(induct u v rule: zeros.induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3312
    fix ia ja
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3313
    assume h: "\<not> ja < ia \<Longrightarrow>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3314
             \<lbrace>st i \<and>* ps (ia + 1) \<and>* ones (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>  ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3315
             \<lbrace>st j \<and>* ps (ja + 1) \<and>* zeros (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3316
    show "\<lbrace>st i \<and>* ps ia \<and>* ones ia ja \<and>* zero (ja + 1)\<rbrace> ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3317
           \<lbrace>st j \<and>* ps (ja + 1) \<and>* zeros ia ja \<and>* zero (ja + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3318
    proof(cases "ja < ia")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3319
      case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3320
      thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3321
        by (simp add: ones.simps zeros.simps sep_conj_ac, simp only:sep_conj_cond,
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3322
               intro tm.pre_condI, simp, hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3323
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3324
      case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3325
      note h[OF False, step]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3326
      from False have ones_eq: "ones ia ja = (one ia \<and>* ones (ia + 1) ja)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3327
        by(simp add: ones.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3328
      from False have zeros_eq: "zeros ia ja = (zero ia \<and>* zeros (ia + 1) ja)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3329
        by(simp add: zeros.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3330
      have s1: "\<lbrace>st i \<and>* ps ia \<and>* one ia \<and>* ones (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>  ?body 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3331
                 \<lbrace>st i \<and>* ps (ia + 1) \<and>* zero ia \<and>* ones (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3332
      proof(cases "ja < ia + 1")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3333
        case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3334
        from True False have "ja = ia" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3335
        thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3336
          apply(simp add: ones.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3337
          by (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3338
      next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3339
        case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3340
        from False have "ones (ia + 1) ja = (one (ia + 1) \<and>* ones (ia + 1 + 1) ja)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3341
          by(simp add: ones.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3342
        thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3343
          by (simp, hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3344
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3345
      have s2: "\<lbrace>st i \<and>* ps (ia + 1) \<and>* zero ia \<and>* ones (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3346
                ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3347
                \<lbrace>st j \<and>* ps (ja + 1) \<and>* zero ia \<and>* zeros (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3348
        by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3349
      from tm.sequencing[OF s1 s2] have 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3350
        "\<lbrace>st i \<and>* ps ia \<and>* one ia \<and>* ones (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>  ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3351
        \<lbrace>st j \<and>* ps (ja + 1) \<and>* zero ia \<and>* zeros (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3352
      thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3353
        unfolding ones_eq zeros_eq by(simp add: sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3354
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3355
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3356
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3357
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3358
lemma  hoare_clear_until_zero_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3359
  assumes "u = v" "x = w + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3360
  shows "\<lbrace>st i ** ps u ** ones v w ** zero x\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3361
              i :[clear_until_zero]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3362
        \<lbrace>st j ** ps x ** zeros v w ** zero x\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3363
  by (unfold assms, rule hoare_clear_until_zero)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3364
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3365
definition "shift_left = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3366
            (TL start exit.
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3367
              TLabel start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3368
                 if_zero exit;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3369
                 move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3370
                 write_one;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3371
                 right_until_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3372
                 move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3373
                 write_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3374
                 move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3375
                 move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3376
                 jmp start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3377
              TLabel exit)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3378
           "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3379
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3380
declare ones_simps[simp del]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3381
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3382
lemma hoare_move_left_reps[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3383
  assumes "ks \<noteq> []" "u = v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3384
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3385
    "\<lbrace>st i ** ps u ** reps v w ks\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3386
         i:[move_left]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3387
     \<lbrace>st j ** ps (u - 1) **  reps v w ks\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3388
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3389
  from `ks \<noteq> []` obtain y ys where eq_ks: "ks = y#ys"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3390
    by (metis neq_Nil_conv)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3391
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3392
    apply (unfold assms eq_ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3393
    apply (case_tac ys, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3394
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3395
      have "(ones v (v + int y)) = (one v \<and>* ones (v + 1) (v + int y))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3396
        by (smt ones_step_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3397
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3398
    apply (unfold this, hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3399
    by (simp add:this, hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3400
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3401
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3402
lemma hoare_shift_left_cons:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3403
  assumes h: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3404
  shows "\<lbrace>st i \<and>* ps u \<and>* tm (u - 1) x \<and>* reps u v ks \<and>* zero (v + 1) \<and>* zero (v + 2) \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3405
                                   i:[shift_left]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3406
         \<lbrace>st j \<and>* ps (v + 2) \<and>* reps (u - 1) (v - 1) ks \<and>* zero v \<and>* zero (v + 1) \<and>* zero (v + 2) \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3407
proof(unfold shift_left_def, intro t_hoare_local t_hoare_label, clarify, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3408
      rule t_hoare_label_last, simp+, clarify, prune)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3409
  show " \<lbrace>st i \<and>* ps u \<and>* tm (u - 1) x \<and>* reps u v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3410
             i :[ (if_zero j ;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3411
                   move_left ;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3412
                   write_one ;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3413
                   right_until_zero ;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3414
                   move_left ; write_zero ; 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3415
                   move_right ; move_right ; jmp i) ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3416
         \<lbrace>st j \<and>* ps (v + 2) \<and>* reps (u - 1) (v - 1) ks \<and>* zero v \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3417
    using h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3418
  proof(induct ks arbitrary:i u v x)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3419
    case (Cons k ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3420
    thus ?case 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3421
    proof(cases "ks = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3422
      let ?body = "i :[ (if_zero j ;  move_left ; write_one ; right_until_zero ;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3423
                   move_left ; write_zero ; move_right ; move_right ; jmp i) ]: j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3424
      case True 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3425
      have "\<lbrace>st i \<and>* ps u \<and>* tm (u - 1) x \<and>* (one u \<and>* ones (u + 1) (u + int k)) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3426
                                          zero (u + int k + 1) \<and>* zero (u + int k + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3427
                         ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3428
            \<lbrace>st j \<and>* ps (u + int k + 2) \<and>* (one (u - 1) \<and>* ones u (u - 1 + int k)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3429
                       zero (u + int k) \<and>* zero (u + int k + 1) \<and>* zero (u + int k + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3430
      apply(rule tm.sequencing [where q = "st i \<and>* ps (u + int k + 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3431
                (one (u - 1) \<and>* ones u (u - 1 + int k)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3432
                zero (u + int k) \<and>* zero (u + int k + 1) \<and>* zero (u + int k + 2)"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3433
          apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3434
          apply (rule_tac p = "st j' \<and>* ps (u - 1) \<and>* ones (u - 1) (u + int k) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3435
                                zero (u + int k + 1) \<and>* zero (u + int k + 2)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3436
            in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3437
          apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3438
          my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3439
            have "(ones (u - 1) (u + int k)) = (ones (u - 1) (u + int k - 1) \<and>* one (u + int k))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3440
              by (smt ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3441
          my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3442
          apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3443
          apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3444
          apply (simp add:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3445
          apply (smt ones.simps sep.mult_assoc sep_conj_commuteI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3446
          apply (simp add:sep_conj_ac)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3447
          apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3448
          apply (smt ones.simps sep.mult_left_commute sep_conj_commuteI this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3449
          by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3450
        with True show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3451
        by (simp add:ones_simps, simp only:sep_conj_cond, intro tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3452
    next 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3453
      case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3454
      let ?body = "i :[ (if_zero j ; move_left ; write_one ;right_until_zero ; move_left ; 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3455
                                write_zero ; move_right ; move_right ; jmp i) ]: j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3456
      have "\<lbrace>st i \<and>* ps u \<and>* tm (u - 1) x \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3457
                zero (u + int k + 1) \<and>* reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3458
                ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3459
            \<lbrace>st j \<and>* ps (v + 2) \<and>* one (u - 1) \<and>* ones u (u - 1 + int k) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3460
                        zero (u + int k) \<and>* reps (1 + (u + int k)) (v - 1) ks \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3461
                                              zero v \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3462
        apply (rule tm.sequencing[where q = "st i \<and>* ps (u + int k + 2) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3463
                  zero (u + int k + 1) \<and>* reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3464
                  zero (v + 2) \<and>* one (u - 1) \<and>* ones u (u - 1 + int k) \<and>* zero (u + int k)"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3465
        apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3466
        apply (rule_tac p = "st j' \<and>* ps (u - 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3467
                               ones (u - 1) (u + int k) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3468
                               zero (u + int k + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3469
                               reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3470
            " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3471
        apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3472
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3473
          have "(ones (u - 1) (u + int k)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3474
            (ones (u - 1) (u + int k - 1) \<and>* one (u + int k))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3475
            by (smt ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3476
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3477
        apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3478
        apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3479
        apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3480
        apply (smt ones.simps sep.mult_assoc sep_conj_commuteI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3481
        apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3482
        apply (smt ones.simps this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3483
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3484
          have eq_u: "1 + (u + int k) = u + int k + 1" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3485
          from Cons.hyps[OF `ks \<noteq> []`, of i "u + int k + 2" Bk v, folded zero_def] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3486
          have "\<lbrace>st i \<and>* ps (u + int k + 2) \<and>* zero (u + int k + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3487
                            reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3488
                               ?body
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3489
                      \<lbrace>st j \<and>* ps (v + 2) \<and>*  reps (1 + (u + int k)) (v - 1) ks \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3490
                                                zero v \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3491
          by (simp add:eq_u)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3492
        my_block_end my_note hh[step] = this 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3493
        by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3494
      thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3495
        by (unfold reps_simp3[OF False], auto simp:sep_conj_ac ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3496
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3497
  qed auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3498
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3499
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3500
lemma hoare_shift_left_cons_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3501
  assumes h: "ks \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3502
          "v = u - 1" "w = u" "y = x + 1" "z = x + 2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3503
  shows "\<lbrace>st i \<and>* ps u \<and>* tm v vv \<and>* reps w x ks \<and>* tm y Bk \<and>* tm z Bk\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3504
                                   i:[shift_left]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3505
         \<lbrace>st j \<and>* ps z \<and>* reps v (x - 1) ks \<and>* zero x \<and>* zero y \<and>* zero z \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3506
  by (unfold assms, fold zero_def, rule hoare_shift_left_cons[OF `ks \<noteq> []`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3507
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3508
definition "bone c1 c2 = (TL exit l_one.
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3509
                                if_one l_one;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3510
                                  (c1;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3511
                                   jmp exit);
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3512
                                TLabel l_one;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3513
                                      c2;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3514
                                TLabel exit
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3515
                              )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3516
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3517
lemma hoare_bone_1_out:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3518
  assumes h: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3519
        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3520
                         i:[c1]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3521
                  \<lbrace>st e \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3522
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3523
  shows "\<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3524
              i:[(bone c1 c2)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3525
         \<lbrace>st e \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3526
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3527
apply (unfold bone_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3528
apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3529
apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3530
apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3531
by (rule h)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3532
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3533
lemma hoare_bone_1:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3534
  assumes h: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3535
        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3536
                         i:[c1]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3537
                  \<lbrace>st j \<and>* ps v \<and>* tm v x \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3538
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3539
  shows "\<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3540
              i:[(bone c1 c2)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3541
         \<lbrace>st j \<and>* ps v \<and>* tm v x \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3542
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3543
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3544
  note h[step]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3545
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3546
    apply (unfold bone_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3547
    apply (rule t_hoare_label_last, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3548
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3549
    apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3550
    by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3551
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3552
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3553
lemma hoare_bone_2:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3554
  assumes h: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3555
        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3556
                         i:[c2]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3557
                  \<lbrace>st j \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3558
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3559
  shows "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3560
              i:[(bone c1 c2)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3561
         \<lbrace>st j \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3562
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3563
apply (unfold bone_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3564
apply (rule_tac q = "st la \<and>* ps u \<and>* one u \<and>* p" in tm.sequencing)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3565
apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3566
apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3567
apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3568
apply (subst tassemble_to.simps(2), intro tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3569
apply (subst tassemble_to.simps(4), intro tm.code_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3570
apply (subst tassemble_to.simps(2), intro tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3571
apply (subst tassemble_to.simps(4), simp add:sep_conj_cond, rule tm.code_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3572
by (rule h)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3573
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3574
lemma hoare_bone_2_out:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3575
  assumes h: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3576
        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3577
                         i:[c2]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3578
                  \<lbrace>st e \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3579
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3580
  shows "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3581
              i:[(bone c1 c2)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3582
         \<lbrace>st e \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3583
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3584
apply (unfold bone_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3585
apply (rule_tac q = "st la \<and>* ps u \<and>* one u \<and>* p" in tm.sequencing)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3586
apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3587
apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3588
apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3589
apply (subst tassemble_to.simps(2), intro tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3590
apply (subst tassemble_to.simps(4), rule tm.code_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3591
apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3592
by (rule h)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3593
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3594
definition "bzero c1 c2 = (TL exit l_zero.
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3595
                                if_zero l_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3596
                                  (c1;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3597
                                   jmp exit);
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3598
                                TLabel l_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3599
                                      c2;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3600
                                TLabel exit
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3601
                              )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3602
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3603
lemma hoare_bzero_1:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3604
  assumes h[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3605
        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3606
                         i:[c1]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3607
                 \<lbrace>st j \<and>* ps v \<and>* tm v x \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3608
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3609
  shows "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3610
              i:[(bzero c1 c2)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3611
         \<lbrace>st j \<and>* ps v \<and>* tm v x \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3612
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3613
apply (unfold bzero_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3614
apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3615
apply (rule_tac c = " ((c1 ; jmp l) ; TLabel la ; c2 ; TLabel l)" in t_hoare_label_last, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3616
apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3617
by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3618
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3619
lemma hoare_bzero_1_out:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3620
  assumes h[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3621
        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3622
                         i:[c1]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3623
                 \<lbrace>st e \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3624
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3625
  shows "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3626
              i:[(bzero c1 c2)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3627
         \<lbrace>st e \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3628
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3629
apply (unfold bzero_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3630
apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3631
apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3632
apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3633
by (rule h)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3634
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3635
lemma hoare_bzero_2:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3636
  assumes h: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3637
        "\<And> i j. \<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3638
                         i:[c2]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3639
                 \<lbrace>st j \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3640
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3641
  shows "\<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3642
              i:[(bzero c1 c2)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3643
         \<lbrace>st j \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3644
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3645
  apply (unfold bzero_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3646
  apply (rule_tac q = "st la \<and>* ps u \<and>* zero u \<and>* p" in tm.sequencing)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3647
  apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3648
  apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3649
  apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3650
  apply (subst tassemble_to.simps(2), intro tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3651
  apply (subst tassemble_to.simps(4))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3652
  apply (rule tm.code_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3653
  apply (subst tassemble_to.simps(2))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3654
  apply (rule tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3655
  apply (subst tassemble_to.simps(4), simp add:sep_conj_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3656
  apply (rule tm.code_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3657
  by (rule h)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3658
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3659
lemma hoare_bzero_2_out:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3660
  assumes h: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3661
        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3662
                         i:[c2]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3663
                  \<lbrace>st e \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3664
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3665
  shows "\<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3666
              i:[(bzero c1 c2)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3667
         \<lbrace>st e \<and>* q \<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3668
        "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3669
apply (unfold bzero_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3670
apply (rule_tac q = "st la \<and>* ps u \<and>* zero u \<and>* p" in tm.sequencing)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3671
apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3672
apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3673
apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3674
apply (subst tassemble_to.simps(2), intro tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3675
apply (subst tassemble_to.simps(4), rule tm.code_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3676
apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3677
by (rule h)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3678
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3679
definition "skip_or_set = bone (write_one; move_right; move_right)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3680
                               (right_until_zero; move_right)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3681
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3682
lemma reps_len_split: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3683
  assumes "xs \<noteq> []" "ys \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3684
  shows "reps_len (xs @ ys) = reps_len xs + reps_len ys + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3685
  using assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3686
proof(induct xs arbitrary:ys)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3687
  case (Cons x1 xs1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3688
  show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3689
  proof(cases "xs1 = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3690
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3691
    thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3692
      by (simp add:reps_len_cons[OF `ys \<noteq> []`] reps_len_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3693
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3694
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3695
    hence " xs1 @ ys \<noteq> []" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3696
    thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3697
      apply (simp add:reps_len_cons[OF `xs1@ys \<noteq> []`] reps_len_cons[OF `xs1 \<noteq> []`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3698
      by (simp add: Cons.hyps[OF `xs1 \<noteq> []` `ys \<noteq> []`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3699
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3700
qed auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3701
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3702
lemma hoare_skip_or_set_set:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3703
  "\<lbrace> st i \<and>* ps u \<and>* zero u \<and>* zero (u + 1) \<and>* tm (u + 2) x\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3704
         i:[skip_or_set]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3705
   \<lbrace> st j \<and>* ps (u + 2) \<and>* one u \<and>* zero (u + 1) \<and>* tm (u + 2) x\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3706
  apply(unfold skip_or_set_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3707
  apply(rule_tac q = "st j \<and>* ps (u + 2) \<and>* tm (u + 2) x \<and>* one u \<and>* zero (u + 1)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3708
    in tm.post_weaken)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3709
  apply(rule hoare_bone_1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3710
  apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3711
  by (auto simp:sep_conj_ac, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3712
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3713
lemma hoare_skip_or_set_set_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3714
  assumes "u = v" "w = v + 1" "x = v + 2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3715
  shows "\<lbrace>st i \<and>* ps u \<and>* zero v \<and>* zero w \<and>* tm x xv\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3716
                   i:[skip_or_set]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3717
         \<lbrace>st j \<and>* ps x \<and>* one v \<and>* zero w \<and>* tm x xv\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3718
  by (unfold assms, rule hoare_skip_or_set_set)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3719
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3720
lemma hoare_skip_or_set_skip:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3721
  "\<lbrace> st i \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3722
         i:[skip_or_set]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3723
   \<lbrace> st j \<and>*  ps (v + 2) \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3724
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3725
   show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3726
     apply(unfold skip_or_set_def, unfold reps.simps, simp add:sep_conj_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3727
     apply(rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3728
     apply(rule_tac p = "st i \<and>* ps u \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3729
                             zero (u + int k + 1)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3730
                   in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3731
     apply (rule_tac q = "st j \<and>* ps (u + int k + 2) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3732
                          one u \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3733
              " in tm.post_weaken)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3734
     apply (rule hoare_bone_2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3735
     apply (rule_tac p = " st i \<and>* ps u \<and>* ones u (u + int k) \<and>* zero (u + int k + 1) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3736
       " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3737
     apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3738
     apply (simp add:sep_conj_ac, sep_cancel+, auto simp:sep_conj_ac ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3739
     by (sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3740
 qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3741
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3742
lemma hoare_skip_or_set_skip_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3743
  assumes "u = v" "x = w + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3744
  shows  "\<lbrace> st i \<and>* ps u \<and>* reps v w [k] \<and>* zero x\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3745
                  i:[skip_or_set]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3746
          \<lbrace> st j \<and>*  ps (w + 2) \<and>* reps v w [k] \<and>* zero x\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3747
  by (unfold assms, rule hoare_skip_or_set_skip)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3748
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3749
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3750
definition "if_reps_z e = (move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3751
                              bone (move_left; jmp e) (move_left)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3752
                             )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3753
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3754
lemma hoare_if_reps_z_true:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3755
  assumes h: "k = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3756
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3757
   "\<lbrace>st i \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3758
      i:[if_reps_z e]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3759
    \<lbrace>st e \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3760
  apply (unfold reps.simps, simp add:sep_conj_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3761
  apply (rule tm.pre_condI, simp add:h)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3762
  apply (unfold if_reps_z_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3763
  apply (simp add:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3764
  apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3765
  apply (rule_tac p = "st j' \<and>* ps (u + 1) \<and>* zero (u + 1) \<and>* one u" in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3766
  apply (rule hoare_bone_1_out)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3767
  by (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3768
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3769
lemma hoare_if_reps_z_true_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3770
  assumes "k = 0" "u = v" "x = w + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3771
  shows "\<lbrace>st i \<and>* ps u \<and>* reps v w [k] \<and>* zero x\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3772
                  i:[if_reps_z e]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3773
         \<lbrace>st e \<and>* ps u \<and>* reps v w [k] \<and>* zero x\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3774
  by (unfold assms, rule hoare_if_reps_z_true, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3775
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3776
lemma hoare_if_reps_z_false:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3777
  assumes h: "k \<noteq> 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3778
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3779
   "\<lbrace>st i \<and>* ps u \<and>* reps u v [k]\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3780
      i:[if_reps_z e]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3781
    \<lbrace>st j \<and>* ps u \<and>* reps u v [k]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3782
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3783
  from h obtain k' where "k = Suc k'" by (metis not0_implies_Suc)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3784
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3785
    apply (unfold `k = Suc k'`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3786
    apply (simp add:sep_conj_cond, rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3787
    apply (unfold if_reps_z_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3788
    apply (simp add:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3789
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3790
    apply (rule_tac p = "st j' \<and>* ps (u + 1) \<and>* one (u + 1) \<and>* one u \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3791
                          ones (2 + u) (u + (1 + int k'))" in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3792
    apply (rule_tac hoare_bone_2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3793
    by (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3794
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3795
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3796
lemma hoare_if_reps_z_false_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3797
  assumes h: "k \<noteq> 0" "u = v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3798
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3799
   "\<lbrace>st i \<and>* ps u \<and>* reps v w [k]\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3800
      i:[if_reps_z e]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3801
    \<lbrace>st j \<and>* ps u \<and>* reps v w [k]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3802
  by (unfold assms, rule hoare_if_reps_z_false[OF `k \<noteq> 0`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3803
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3804
definition "if_reps_nz e = (move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3805
                              bzero (move_left; jmp e) (move_left)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3806
                           )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3807
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3808
lemma EXS_postI: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3809
  assumes "\<lbrace>P\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3810
            c
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3811
           \<lbrace>Q x\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3812
  shows "\<lbrace>P\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3813
          c
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3814
        \<lbrace>EXS x. Q x\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3815
by (metis EXS_intro assms tm.hoare_adjust)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3816
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3817
lemma hoare_if_reps_nz_true:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3818
  assumes h: "k \<noteq> 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3819
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3820
   "\<lbrace>st i \<and>* ps u \<and>* reps u v [k]\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3821
      i:[if_reps_nz e]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3822
    \<lbrace>st e \<and>* ps u \<and>* reps u v [k]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3823
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3824
  from h obtain k' where "k = Suc k'" by (metis not0_implies_Suc)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3825
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3826
    apply (unfold `k = Suc k'`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3827
    apply (unfold reps.simps, simp add:sep_conj_cond, rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3828
    apply (unfold if_reps_nz_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3829
    apply (simp add:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3830
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3831
    apply (rule_tac p = "st j' \<and>* ps (u + 1) \<and>* one (u + 1) \<and>* one u \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3832
                            ones (2 + u) (u + (1 + int k'))" in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3833
    apply (rule hoare_bzero_1_out)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3834
    by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3835
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3836
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3837
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3838
lemma hoare_if_reps_nz_true_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3839
  assumes h: "k \<noteq> 0" "u = v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3840
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3841
   "\<lbrace>st i \<and>* ps u \<and>* reps v w [k]\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3842
      i:[if_reps_nz e]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3843
    \<lbrace>st e \<and>* ps u \<and>* reps v w [k]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3844
  by (unfold assms, rule hoare_if_reps_nz_true[OF `k\<noteq> 0`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3845
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3846
lemma hoare_if_reps_nz_false:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3847
  assumes h: "k = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3848
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3849
   "\<lbrace>st i \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3850
      i:[if_reps_nz e]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3851
    \<lbrace>st j \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3852
  apply (simp add:h sep_conj_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3853
  apply (rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3854
  apply (unfold if_reps_nz_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3855
  apply (simp add:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3856
  apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3857
  apply (rule_tac p = "st j' \<and>* ps (u + 1) \<and>*  zero (u + 1) \<and>* one u" in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3858
  apply (rule hoare_bzero_2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3859
  by (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3860
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3861
lemma hoare_if_reps_nz_false_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3862
  assumes h: "k = 0" "u = v" "x = w + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3863
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3864
   "\<lbrace>st i \<and>* ps u \<and>* reps v w [k] \<and>* zero x\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3865
      i:[if_reps_nz e]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3866
    \<lbrace>st j \<and>* ps u \<and>* reps v w [k] \<and>* zero x\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3867
  by (unfold assms, rule hoare_if_reps_nz_false, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3868
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3869
definition "skip_or_sets n = tpg_fold (replicate n skip_or_set)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3870
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3871
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3872
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3873
lemma hoare_skip_or_sets_set:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3874
  shows "\<lbrace>st i \<and>* ps u \<and>* zeros u (u + int (reps_len (replicate (Suc n) 0))) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3875
                                  tm (u + int (reps_len (replicate (Suc n) 0)) + 1) x\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3876
            i:[skip_or_sets (Suc n)]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3877
         \<lbrace>st j \<and>* ps (u + int (reps_len (replicate (Suc n) 0)) + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3878
                     reps' u  (u + int (reps_len (replicate (Suc n) 0))) (replicate (Suc n) 0) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3879
                                 tm (u + int (reps_len (replicate (Suc n) 0)) + 1) x \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3880
proof(induct n arbitrary:i j u x)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3881
  case 0
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3882
  from 0 show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3883
    apply (simp add:reps'_def reps_len_def reps_ctnt_len_def reps_sep_len_def reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3884
    apply (unfold skip_or_sets_def, simp add:tpg_fold_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3885
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3886
    by (auto simp:sep_conj_ac, smt cond_true_eq2 ones.simps sep_conj_left_commute)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3887
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3888
    case (Suc n)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3889
    { fix n
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3890
      have "listsum (replicate n (Suc 0)) = n"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3891
        by (induct n, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3892
    } note eq_sum = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3893
    have eq_len: "\<And>n. n \<noteq> 0 \<Longrightarrow> reps_len (replicate (Suc n) 0) = reps_len (replicate n 0) + 2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3894
      by (simp add:reps_len_def reps_sep_len_def reps_ctnt_len_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3895
    have eq_zero: "\<And> u v. (zeros u (u + int (v + 2))) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3896
           (zeros u (u + (int v)) \<and>* zero (u + (int v) + 1) \<and>* zero (u + (int v) + 2))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3897
      by (smt sep.mult_assoc zeros_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3898
    hence eq_z: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3899
      "zeros u (u + int (reps_len (replicate (Suc (Suc n)) 0)))  = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3900
       (zeros u (u + int (reps_len (replicate (Suc n) 0))) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3901
       zero ((u + int (reps_len (replicate (Suc n) 0))) + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3902
       zero ((u + int (reps_len (replicate (Suc n) 0))) + 2))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3903
      " by (simp only:eq_len)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3904
    have hh: "\<And>x. (replicate (Suc (Suc n)) x) = (replicate (Suc n) x) @ [x]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3905
      by (metis replicate_Suc replicate_append_same)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3906
    have hhh: "replicate (Suc n) skip_or_set \<noteq> []" "[skip_or_set] \<noteq> []" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3907
    have eq_code: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3908
          "(i :[ skip_or_sets (Suc (Suc n)) ]: j) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3909
           (i :[ (skip_or_sets (Suc n); skip_or_set) ]: j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3910
    proof(unfold skip_or_sets_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3911
      show "i :[ tpg_fold (replicate (Suc (Suc n)) skip_or_set) ]: j =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3912
               i :[ (tpg_fold (replicate (Suc n) skip_or_set) ; skip_or_set) ]: j"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3913
        apply (insert tpg_fold_app[OF hhh, of i j], unfold hh)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3914
        by (simp only:tpg_fold_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3915
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3916
    have "Suc n \<noteq> 0" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3917
    show ?case 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3918
      apply (unfold eq_z eq_code)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3919
      apply (hstep Suc(1))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3920
      apply (unfold eq_len[OF `Suc n \<noteq> 0`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3921
      apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3922
      apply (auto simp:sep_conj_ac)[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3923
      apply (sep_cancel+, prune) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3924
      apply (fwd abs_ones)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3925
      apply ((fwd abs_reps')+, simp add:int_add_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3926
      by (metis replicate_append_same)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3927
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3928
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3929
lemma hoare_skip_or_sets_set_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3930
  assumes h: "p2 = p1" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3931
             "p3 = p1 + int (reps_len (replicate (Suc n) 0))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3932
             "p4 = p3 + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3933
  shows "\<lbrace>st i \<and>* ps p1 \<and>* zeros p2 p3 \<and>* tm p4 x\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3934
            i:[skip_or_sets (Suc n)]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3935
         \<lbrace>st j \<and>* ps p4 \<and>* reps' p2  p3 (replicate (Suc n) 0) \<and>* tm p4 x\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3936
  apply (unfold h)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3937
  by (rule hoare_skip_or_sets_set)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3938
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3939
declare reps.simps[simp del]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3940
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3941
lemma hoare_skip_or_sets_skip:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3942
  assumes h: "n < length ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3943
  shows "\<lbrace>st i \<and>* ps u \<and>* reps' u v (take n ks) \<and>* reps' (v + 1) w [ks!n] \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3944
            i:[skip_or_sets (Suc n)]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3945
         \<lbrace>st j \<and>* ps (w+1) \<and>* reps' u v (take n ks) \<and>* reps' (v + 1) w [ks!n]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3946
  using h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3947
proof(induct n arbitrary: i j u v w ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3948
  case 0
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3949
  show ?case 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3950
    apply (subst (1 5) reps'_def, simp add:sep_conj_cond, intro tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3951
    apply (unfold skip_or_sets_def, simp add:tpg_fold_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3952
    apply (unfold reps'_def, simp del:reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3953
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3954
    by (sep_cancel+, smt+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3955
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3956
  case (Suc n)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3957
  from `Suc n < length ks` have "n < length ks" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3958
  note h =  Suc(1) [OF this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3959
  show ?case 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3960
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3961
      from `Suc n < length ks` 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3962
      have eq_take: "take (Suc n) ks = take n ks @ [ks!n]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3963
        by (metis not_less_eq not_less_iff_gr_or_eq take_Suc_conv_app_nth)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3964
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3965
    apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3966
    apply (subst reps'_append, simp add:sep_conj_exists, rule tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3967
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3968
      have "(i :[ skip_or_sets (Suc (Suc n)) ]: j) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3969
                 (i :[ (skip_or_sets (Suc n); skip_or_set )]: j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3970
      proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3971
        have eq_rep: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3972
          "(replicate (Suc (Suc n)) skip_or_set) = ((replicate (Suc n) skip_or_set) @ [skip_or_set])"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3973
          by (metis replicate_Suc replicate_append_same)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3974
        have "replicate (Suc n) skip_or_set \<noteq> []" "[skip_or_set] \<noteq> []" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3975
        from tpg_fold_app[OF this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3976
        show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3977
          by (unfold skip_or_sets_def eq_rep, simp del:replicate.simps add:tpg_fold_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3978
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3979
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3980
    apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3981
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3982
       fix i j m 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3983
       have "\<lbrace>st i \<and>* ps u \<and>* (reps' u (m - 1) (take n ks) \<and>* reps' m v [ks ! n])\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3984
                            i :[ (skip_or_sets (Suc n)) ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3985
             \<lbrace>st j \<and>* ps (v + 1) \<and>* (reps' u (m - 1) (take n ks) \<and>* reps' m v [ks ! n])\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3986
                  apply (rule h[THEN tm.hoare_adjust])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3987
                  by (sep_cancel+, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3988
    my_block_end my_note h_c1 = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3989
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3990
      fix j' j m 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3991
      have "\<lbrace>st j' \<and>* ps (v + 1) \<and>* reps' (v + 1) w [ks ! Suc n]\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3992
                          j' :[ skip_or_set ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3993
            \<lbrace>st j \<and>* ps (w + 1) \<and>* reps' (v + 1) w [ks ! Suc n]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3994
        apply (unfold reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3995
        apply (rule hoare_skip_or_set_skip[THEN tm.hoare_adjust])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3996
        by (sep_cancel+, smt)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3997
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3998
    apply (hstep h_c1 this)+ 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  3999
    by ((fwd abs_reps'), simp, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4000
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4001
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4002
lemma hoare_skip_or_sets_skip_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4003
  assumes h: "n < length ks" "u = v" "x = w + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4004
  shows "\<lbrace>st i \<and>* ps u \<and>* reps' v w (take n ks) \<and>* reps' x y [ks!n] \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4005
            i:[skip_or_sets (Suc n)]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4006
         \<lbrace>st j \<and>* ps (y+1) \<and>* reps' v w (take n ks) \<and>* reps' x y [ks!n]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4007
  by (unfold assms, rule hoare_skip_or_sets_skip[OF `n < length ks`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4008
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4009
lemma fam_conj_interv_simp:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4010
    "(fam_conj {(ia::int)<..} p) = ((p (ia + 1)) \<and>* fam_conj {ia + 1 <..} p)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4011
by (smt Collect_cong fam_conj_insert_simp greaterThan_def 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4012
        greaterThan_eq_iff greaterThan_iff insertI1 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4013
        insert_compr lessThan_iff mem_Collect_eq)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4014
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4015
lemma zeros_fam_conj:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4016
  assumes "u \<le> v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4017
  shows "(zeros u v \<and>* fam_conj {v<..} zero) = fam_conj {u - 1<..} zero"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4018
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4019
  have "{u - 1<..v} ## {v <..}" by (auto simp:set_ins_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4020
  from fam_conj_disj_simp[OF this, symmetric]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4021
  have "(fam_conj {u - 1<..v} zero \<and>* fam_conj {v<..} zero) = fam_conj ({u - 1<..v} + {v<..}) zero" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4022
  moreover 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4023
  from `u \<le> v` have eq_set: "{u - 1 <..} = {u - 1 <..v} + {v <..}" by (auto simp:set_ins_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4024
  moreover have "fam_conj {u - 1<..v} zero = zeros u v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4025
  proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4026
    have "({u - 1<..v}) = ({u .. v})" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4027
    moreover {
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4028
      fix u v 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4029
      assume "u  \<le> (v::int)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4030
      hence "fam_conj {u .. v} zero = zeros u v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4031
      proof(induct rule:ones_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4032
        case (Base i j)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4033
        thus ?case by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4034
      next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4035
        case (Step i j)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4036
        thus ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4037
        proof(cases "i = j") 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4038
          case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4039
          show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4040
            by (unfold True, simp add:fam_conj_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4041
        next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4042
          case False 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4043
          with `i \<le> j` have hh: "i + 1 \<le> j" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4044
          hence eq_set: "{i..j} = (insert i {i + 1 .. j})"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4045
            by (smt simp_from_to)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4046
          have "i \<notin> {i + 1 .. j}" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4047
          from fam_conj_insert_simp[OF this, folded eq_set]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4048
          have "fam_conj {i..j} zero = (zero i \<and>* fam_conj {i + 1..j} zero)" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4049
          with Step(2)[OF hh] Step
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4050
          show ?thesis by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4051
        qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4052
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4053
    } 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4054
    moreover note this[OF `u  \<le> v`]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4055
    ultimately show ?thesis by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4056
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4057
  ultimately show ?thesis by smt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4058
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4059
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4060
declare replicate.simps [simp del]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4061
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4062
lemma hoare_skip_or_sets_comb:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4063
  assumes "length ks \<le> n"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4064
  shows "\<lbrace>st i \<and>* ps u \<and>* reps u v ks \<and>* fam_conj {v<..} zero\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4065
                i:[skip_or_sets (Suc n)]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4066
         \<lbrace>st j \<and>* ps ((v + int (reps_len (list_ext n ks)) - int (reps_len ks))+ 2) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4067
          reps' u (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1) (list_ext n ks) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4068
          fam_conj {(v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1)<..} zero \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4069
proof(cases "ks = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4070
  case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4071
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4072
    apply (subst True, simp only:reps.simps sep_conj_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4073
    apply (rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4074
    apply (rule_tac p = "st i \<and>* ps (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4075
            zeros (v + 1) (v + 1 + int (reps_len (replicate (Suc n) 0))) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4076
            tm (v + 2 + int (reps_len (replicate (Suc n) 0))) Bk \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4077
            fam_conj {(v + 2 + int (reps_len (replicate (Suc n) 0)))<..} zero
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4078
      " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4079
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4080
    apply (auto simp:sep_conj_ac)[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4081
    apply (auto simp:sep_conj_ac)[2]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4082
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4083
      from True have "(list_ext n ks) = (replicate (Suc n) 0)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4084
        by (metis append_Nil diff_zero list.size(3) list_ext_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4085
    my_block_end my_note le_red = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4086
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4087
      from True have "(reps_len ks) = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4088
        by (metis reps_len_nil)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4089
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4090
    apply (unfold this le_red, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4091
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4092
      have "v + 2 + int (reps_len (replicate (Suc n) 0)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4093
            v + int (reps_len (replicate (Suc n) 0)) + 2" by smt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4094
    my_block_end my_note eq_len = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4095
    apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4096
    apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4097
    apply (fold zero_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4098
    apply (subst fam_conj_interv_simp, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4099
    apply (simp only:int_add_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4100
    apply (simp only:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4101
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4102
      have "v + 1 \<le> (2 + (v + int (reps_len (replicate (Suc n) 0))))" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4103
      from zeros_fam_conj[OF this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4104
      have "(fam_conj {v<..} zero) = (zeros (v + 1) (2 + (v + int (reps_len (replicate (Suc n) 0)))) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4105
                                        fam_conj {2 + (v + int (reps_len (replicate (Suc n) 0)))<..} zero)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4106
        by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4107
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4108
    apply (subst (asm) this, simp only:int_add_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4109
    by (smt cond_true_eq2 sep.mult_assoc sep.mult_commute 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4110
            sep.mult_left_commute sep_conj_assoc sep_conj_commute 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4111
         sep_conj_left_commute zeros.simps zeros_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4112
next 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4113
  case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4114
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4115
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4116
      have "(i:[skip_or_sets (Suc n)]:j) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4117
              (i:[(skip_or_sets (length ks);  skip_or_sets (Suc n - length ks))]:j)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4118
        apply (unfold skip_or_sets_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4119
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4120
          have "(replicate (Suc n) skip_or_set) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4121
                   (replicate (length ks) skip_or_set @ (replicate (Suc n - length ks) skip_or_set))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4122
            by (smt assms replicate_add)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4123
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4124
        apply (unfold this, rule tpg_fold_app, simp add:False)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4125
        by (insert `length ks \<le> n`, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4126
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4127
    apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4128
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4129
      from False have "length ks = (Suc (length ks - 1))" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4130
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4131
    apply (subst (1) this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4132
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4133
      from False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4134
      have "(reps u v ks \<and>* fam_conj {v<..} zero) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4135
            (reps' u (v + 1) ks \<and>* fam_conj {v+1<..} zero)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4136
        apply (unfold reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4137
        by (subst fam_conj_interv_simp, simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4138
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4139
    apply (unfold this) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4140
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4141
      fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4142
      have "\<lbrace>st i \<and>* ps u \<and>* reps' u (v + 1) ks \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4143
                i :[ skip_or_sets (Suc (length ks - 1))]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4144
            \<lbrace>st j \<and>* ps (v + 2) \<and>* reps' u (v + 1) ks \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4145
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4146
          have "ks = take (length ks - 1) ks @ [ks!(length ks - 1)]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4147
            by (smt False drop_0 drop_eq_Nil id_take_nth_drop)  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4148
        my_block_end my_note eq_ks = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4149
        apply (subst (1) this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4150
        apply (unfold reps'_append, simp add:sep_conj_exists, rule tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4151
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4152
          from False have "(length ks - Suc 0) < length ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4153
            by (smt `length ks = Suc (length ks - 1)`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4154
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4155
        apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4156
        apply (subst eq_ks, unfold reps'_append, simp only:sep_conj_exists)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4157
        by (rule_tac x = m in EXS_intro, simp add:sep_conj_ac, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4158
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4159
    apply (hstep this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4160
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4161
      fix u n
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4162
      have "(fam_conj {u <..} zero) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4163
         (zeros (u + 1) (u + int n + 1) \<and>* tm (u + int n + 2) Bk \<and>* fam_conj {(u + int n + 2)<..} zero)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4164
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4165
          have "u + 1 \<le> (u + int n + 2)" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4166
          from zeros_fam_conj[OF this, symmetric]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4167
          have "fam_conj {u<..} zero = (zeros (u + 1) (u + int n + 2) \<and>* fam_conj {u + int n + 2<..} zero)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4168
            by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4169
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4170
        apply (subst this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4171
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4172
          have "(zeros (u + 1) (u + int n + 2)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4173
                   ((zeros (u + 1) (u + int n + 1) \<and>* tm (u + int n + 2) Bk))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4174
            by (smt zero_def zeros_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4175
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4176
        by (unfold this, auto simp:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4177
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4178
    apply (subst (1) this[of _ "(reps_len (replicate (Suc (n - length ks)) 0))"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4179
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4180
      from `length ks \<le> n`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4181
      have "Suc n - length ks = Suc (n - length ks)" by auto 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4182
    my_block_end my_note eq_suc = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4183
    apply (subst this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4184
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4185
    apply (simp add: sep_conj_ac this, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4186
    apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4187
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4188
      have "(int (reps_len (replicate (Suc (n - length ks)) 0))) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4189
            (int (reps_len (list_ext n ks)) - int (reps_len ks) - 1)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4190
        apply (unfold list_ext_def eq_suc)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4191
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4192
          have "replicate (Suc (n - length ks)) 0 \<noteq> []" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4193
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4194
        by (unfold reps_len_split[OF False this], simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4195
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4196
    apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4197
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4198
      from `length ks \<le> n`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4199
      have "(ks @ replicate (Suc (n - length ks)) 0) =  (list_ext n ks)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4200
        by (unfold list_ext_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4201
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4202
    apply (unfold this, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4203
    apply (subst fam_conj_interv_simp, unfold zero_def, simp, simp add:int_add_ac sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4204
    by (sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4205
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4206
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4207
lemma hoare_skip_or_sets_comb_gen:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4208
  assumes "length ks \<le> n" "u = v" "w = x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4209
  shows "\<lbrace>st i \<and>* ps u \<and>* reps v w ks \<and>* fam_conj {x<..} zero\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4210
                i:[skip_or_sets (Suc n)]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4211
         \<lbrace>st j \<and>* ps ((x + int (reps_len (list_ext n ks)) - int (reps_len ks))+ 2) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4212
          reps' u (x + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1) (list_ext n ks) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4213
          fam_conj {(x + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1)<..} zero \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4214
  by (unfold assms, rule hoare_skip_or_sets_comb[OF `length ks \<le> n`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4215
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4216
definition "locate n = (skip_or_sets (Suc n);
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4217
                        move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4218
                        move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4219
                        left_until_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4220
                        move_right
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4221
                       )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4222
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4223
lemma list_ext_tail_expand:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4224
  assumes h: "length ks \<le> a"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4225
  shows "list_ext a ks = take a (list_ext a ks) @ [(list_ext a ks)!a]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4226
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4227
  let ?l = "list_ext a ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4228
  from h have eq_len: "length ?l = Suc a"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4229
    by (smt list_ext_len_eq)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4230
  hence "?l \<noteq> []" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4231
  hence "?l = take (length ?l - 1) ?l @ [?l!(length ?l - 1)]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4232
    by (metis `length (list_ext a ks) = Suc a` diff_Suc_1 le_refl 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4233
                    lessI take_Suc_conv_app_nth take_all)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4234
  from this[unfolded eq_len]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4235
  show ?thesis by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4236
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4237
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4238
lemma reps'_nn_expand:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4239
  assumes "xs \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4240
  shows "(reps' u v xs) = (reps u (v - 1) xs \<and>* zero v)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4241
  by (metis assms reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4242
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4243
lemma sep_conj_st1: "(p \<and>* st t \<and>* q) = (st t \<and>* p \<and>* q)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4244
  by (simp only:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4245
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4246
lemma sep_conj_st2: "(p \<and>* st t) = (st t \<and>* p)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4247
  by (simp only:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4248
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4249
lemma sep_conj_st3: "((st t \<and>* p) \<and>* r) = (st t \<and>* p \<and>* r)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4250
  by (simp only:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4251
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4252
lemma sep_conj_st4: "(EXS x. (st t \<and>* r x)) = ((st t) \<and>* (EXS x. r x))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4253
  apply (unfold pred_ex_def, default+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4254
  apply (safe)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4255
  apply (sep_cancel, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4256
  by (auto elim!: sep_conjE intro!:sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4257
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4258
lemmas sep_conj_st = sep_conj_st1 sep_conj_st2 sep_conj_st3 sep_conj_st4
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4259
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4260
lemma sep_conj_cond3 : "(<cond1> \<and>* <cond2>) = <(cond1 \<and> cond2)>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4261
  by (smt cond_eqI cond_true_eq sep_conj_commute sep_conj_empty)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4262
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4263
lemma sep_conj_cond4 : "(<cond1> \<and>* <cond2> \<and>* r) = (<(cond1 \<and> cond2)> \<and>* r)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4264
  by (metis Hoare_tm3.sep_conj_cond3 sep_conj_assoc)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4265
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4266
lemmas sep_conj_cond = sep_conj_cond3 sep_conj_cond4 sep_conj_cond 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4267
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4268
lemma hoare_left_until_zero_reps: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4269
  "\<lbrace>st i ** ps v ** zero u ** reps (u + 1) v [k]\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4270
        i:[left_until_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4271
   \<lbrace>st j ** ps u ** zero u ** reps (u + 1) v [k]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4272
  apply (unfold reps.simps, simp only:sep_conj_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4273
  apply (rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4274
  by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4275
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4276
lemma hoare_left_until_zero_reps_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4277
  assumes "u = x" "w = v + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4278
  shows "\<lbrace>st i ** ps u ** zero v ** reps w x [k]\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4279
                i:[left_until_zero]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4280
         \<lbrace>st j ** ps v ** zero v ** reps w x [k]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4281
  by (unfold assms, rule hoare_left_until_zero_reps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4282
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4283
lemma reps_lenE:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4284
  assumes "reps u v ks s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4285
  shows "( <(v = u + int (reps_len ks) - 1)> \<and>* reps u v ks ) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4286
proof(rule condI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4287
  from reps_len_correct[OF assms] show "v = u + int (reps_len ks) - 1" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4288
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4289
  from assms show "reps u v ks s" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4290
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4291
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4292
lemma condI1: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4293
  assumes "p s" "b"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4294
  shows "(<b> \<and>* p) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4295
proof(rule condI[OF assms(2)])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4296
  from  assms(1) show "p s" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4297
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4298
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4299
lemma hoare_locate_set:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4300
  assumes "length ks \<le> n"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4301
  shows "\<lbrace>st i \<and>* zero (u - 1) \<and>* ps u \<and>* reps u v ks \<and>* fam_conj {v<..} zero\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4302
                i:[locate n]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4303
         \<lbrace>st j \<and>* zero (u - 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4304
             (EXS m w. ps m \<and>* reps' u (m - 1) (take n (list_ext n ks)) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4305
                         reps m w [(list_ext n ks)!n] \<and>* fam_conj {w<..} zero)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4306
proof(cases "take n (list_ext n ks) = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4307
  case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4308
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4309
    apply (unfold locate_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4310
    apply (hstep hoare_skip_or_sets_comb_gen)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4311
    apply (subst (3) list_ext_tail_expand[OF `length ks \<le> n`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4312
    apply (subst (1) reps'_append, simp add:sep_conj_exists)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4313
    apply (rule tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4314
    apply (subst (1) reps'_nn_expand[OF False]) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4315
    apply (rule_tac p = "st j' \<and>* <(m = u + int (reps_len (take n (list_ext n ks))) + 1)> \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4316
            ps (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4317
            ((reps u (m - 1 - 1) (take n (list_ext n ks)) \<and>* zero (m - 1)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4318
             reps' m (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4319
              [list_ext n ks ! n]) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4320
            fam_conj
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4321
             {v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1<..}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4322
             zero \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4323
            zero (u - 1)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4324
      in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4325
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4326
      have "[list_ext n ks ! n] \<noteq> []" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4327
      from reps'_nn_expand[OF this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4328
      have "(reps' m (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1) [list_ext n ks ! n]) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4329
                (reps m (v + (int (reps_len (list_ext n ks)) - int (reps_len ks))) [list_ext n ks ! n] \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4330
                   zero (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1))" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4331
        by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4332
    my_block_end 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4333
    apply (subst this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4334
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4335
      have "(fam_conj {v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1<..} zero) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4336
             (zero (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 2) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4337
              fam_conj {v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 2<..} zero)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4338
        by (subst fam_conj_interv_simp, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4339
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4340
    apply (unfold this) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4341
    apply (simp only:sep_conj_st)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4342
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4343
    apply (auto simp:sep_conj_ac)[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4344
    apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4345
    apply (rule_tac x = m in EXS_intro)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4346
    apply (rule_tac x = "m + int (list_ext n ks ! n)" in EXS_intro)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4347
    apply (simp add:sep_conj_ac del:ones_simps, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4348
    apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4349
    apply (erule_tac condE, clarsimp, simp add:sep_conj_ac int_add_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4350
    apply (fwd abs_reps')
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4351
    apply (fwd abs_reps')
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4352
    apply (simp add:sep_conj_ac int_add_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4353
    apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4354
    apply (subst (asm) reps'_def, subst fam_conj_interv_simp, subst fam_conj_interv_simp, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4355
           simp add:sep_conj_ac int_add_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4356
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4357
      fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4358
      assume h: "(reps (1 + (u + int (reps_len (take n (list_ext n ks)))))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4359
             (v + (- int (reps_len ks) + int (reps_len (list_ext n ks)))) [list_ext n ks ! n]) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4360
      (is "?P s")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4361
      from reps_len_correct[OF this] list_ext_tail_expand[OF `length ks \<le> n`]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4362
      have hh: "v + (- int (reps_len ks) + 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4363
                    int (reps_len (take n (list_ext n ks) @ [list_ext n ks ! n]))) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4364
                  1 + (u + int (reps_len (take n (list_ext n ks)))) + 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4365
                       int (reps_len [list_ext n ks ! n]) - 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4366
        by metis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4367
      have "[list_ext n ks ! n] \<noteq> []" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4368
      from hh[unfolded reps_len_split[OF False this]]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4369
      have "v  =  u + (int (reps_len ks)) - 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4370
        by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4371
      from condI1[where p = ?P, OF h this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4372
      have "(<(v = u + int (reps_len ks) - 1)> \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4373
             reps (1 + (u + int (reps_len (take n (list_ext n ks)))))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4374
             (v + (- int (reps_len ks) + int (reps_len (list_ext n ks)))) [list_ext n ks ! n]) s" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4375
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4376
    apply (fwd this, (subst (asm) sep_conj_cond)+, erule condE, simp add:sep_conj_ac int_add_ac
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4377
              reps_len_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4378
    apply (fwd reps_lenE, (subst (asm) sep_conj_cond)+, erule condE, simp add:sep_conj_ac int_add_ac
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4379
            reps_len_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4380
    by (fwd reps_lenE, (subst (asm) sep_conj_cond)+, erule condE, simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4381
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4382
  case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4383
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4384
    apply (unfold locate_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4385
    apply (hstep hoare_skip_or_sets_comb)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4386
    apply (subst (3) list_ext_tail_expand[OF `length ks \<le> n`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4387
    apply (subst (1) reps'_append, simp add:sep_conj_exists)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4388
    apply (rule tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4389
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4390
      from True `length ks \<le> n`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4391
      have "ks = []" "n = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4392
        apply (metis le0 le_antisym length_0_conv less_nat_zero_code list_ext_len take_eq_Nil)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4393
        by (smt True length_take list.size(3) list_ext_len)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4394
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4395
    apply (unfold True this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4396
    apply (simp add:reps'_def list_ext_def reps.simps reps_len_def reps_sep_len_def 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4397
                 reps_ctnt_len_def
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4398
      del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4399
    apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4400
    apply (rule tm.pre_condI, simp del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4401
    apply (subst fam_conj_interv_simp, simp add:sep_conj_st del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4402
    apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4403
    apply (auto simp:sep_conj_ac)[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4404
    apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4405
    apply (rule_tac x = "(v + int (listsum (replicate (Suc 0) (Suc 0))))" in EXS_intro)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4406
    apply (simp only:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4407
    apply (auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4408
    apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4409
    apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4410
    by smt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4411
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4412
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4413
lemma hoare_locate_set_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4414
  assumes "length ks \<le> n" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4415
           "u = v - 1" "v = w" "x = y"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4416
  shows "\<lbrace>st i \<and>* zero u \<and>* ps v \<and>* reps w x ks \<and>* fam_conj {y<..} zero\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4417
                i:[locate n]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4418
         \<lbrace>st j \<and>* zero u \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4419
             (EXS m w. ps m \<and>* reps' v (m - 1) (take n (list_ext n ks)) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4420
                         reps m w [(list_ext n ks)!n] \<and>* fam_conj {w<..} zero)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4421
  by (unfold assms, rule hoare_locate_set[OF `length ks \<le> n`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4422
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4423
lemma hoare_locate_skip: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4424
  assumes h: "n < length ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4425
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4426
   "\<lbrace>st i \<and>* ps u \<and>* zero (u - 1) \<and>* reps' u (v - 1) (take n ks) \<and>* reps' v w [ks!n] \<and>* tm (w + 1) x\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4427
      i:[locate n]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4428
    \<lbrace>st j \<and>* ps v \<and>* zero (u - 1) \<and>* reps' u (v - 1) (take n ks) \<and>* reps' v w [ks!n] \<and>* tm (w + 1) x\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4429
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4430
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4431
    apply (unfold locate_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4432
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4433
    apply (subst (2 4) reps'_def, simp add:reps.simps sep_conj_cond del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4434
    apply (intro tm.pre_condI, simp del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4435
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4436
    apply (case_tac "(take n ks) = []", simp add:reps'_def sep_conj_cond del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4437
    apply (rule tm.pre_condI, simp del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4438
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4439
    apply (simp del:ones_simps add:reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4440
    by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4441
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4442
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4443
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4444
lemma hoare_locate_skip_gen[step]: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4445
  assumes "n < length ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4446
          "v = u - 1" "w = u" "x = y - 1" "z' = z + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4447
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4448
   "\<lbrace>st i \<and>* ps u \<and>* tm v Bk \<and>* reps' w x (take n ks) \<and>* reps' y z [ks!n] \<and>* tm z' vx\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4449
      i:[locate n]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4450
    \<lbrace>st j \<and>* ps y \<and>* tm v Bk \<and>* reps' w x (take n ks) \<and>* reps' y z [ks!n] \<and>* tm z' vx\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4451
  by (unfold assms, fold zero_def, rule hoare_locate_skip[OF `n < length ks`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4452
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4453
definition "Inc a = locate a; 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4454
                    right_until_zero; 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4455
                    move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4456
                    shift_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4457
                    move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4458
                    left_until_double_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4459
                    write_one;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4460
                    left_until_double_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4461
                    move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4462
                    move_right
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4463
                    "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4464
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4465
lemma ones_int_expand: "(ones m (m + int k)) = (one m \<and>* ones (m + 1) (m + int k))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4466
  by (simp add:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4467
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4468
lemma reps_splitedI:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4469
  assumes h1: "(reps u v ks1 \<and>* zero (v + 1) \<and>* reps (v + 2) w ks2) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4470
  and h2: "ks1 \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4471
  and h3: "ks2 \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4472
  shows "(reps u w (ks1 @ ks2)) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4473
proof - 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4474
  from h2 h3
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4475
  have "splited (ks1 @ ks2) ks1 ks2" by (auto simp:splited_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4476
  from h1 obtain s1 where 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4477
    "(reps u v ks1) s1" by (auto elim:sep_conjE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4478
  from h1 obtain s2 where 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4479
    "(reps (v + 2) w ks2) s2" by (auto elim:sep_conjE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4480
  from reps_len_correct[OF `(reps u v ks1) s1`] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4481
  have eq_v: "v = u + int (reps_len ks1) - 1" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4482
  from reps_len_correct[OF `(reps (v + 2) w ks2) s2`]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4483
  have eq_w: "w = v + 2 + int (reps_len ks2) - 1" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4484
  from h1
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4485
  have "(reps u (u + int (reps_len ks1) - 1) ks1 \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4486
         zero (u + int (reps_len ks1)) \<and>* reps (u + int (reps_len ks1) + 1) w ks2) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4487
    apply (unfold eq_v eq_w[unfolded eq_v])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4488
    by (sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4489
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4490
    by(unfold reps_splited[OF `splited (ks1 @ ks2) ks1 ks2`], simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4491
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4492
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4493
lemma reps_sucI:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4494
  assumes h: "(reps u v (xs@[x]) \<and>* one (1 + v)) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4495
  shows "(reps u (1 + v) (xs@[Suc x])) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4496
proof(cases "xs = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4497
  case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4498
  from h obtain s' where "(reps u v (xs@[x])) s'" by (auto elim:sep_conjE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4499
  from reps_len_correct[OF this] have " v = u + int (reps_len (xs @ [x])) - 1" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4500
  with True have eq_v: "v = u + int x" by (simp add:reps_len_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4501
  have eq_one1: "(one (1 + (u + int x)) \<and>* ones (u + 1) (u + int x)) = ones (u + 1) (1 + (u + int x))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4502
    by (smt ones_rev sep.mult_commute)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4503
  from h show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4504
    apply (unfold True, simp add:eq_v reps.simps sep_conj_cond sep_conj_ac ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4505
    by (sep_cancel+, simp add: eq_one1, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4506
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4507
  case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4508
  from h obtain s1 s2 where hh: "(reps u v (xs@[x])) s1" "s = s1 + s2" "s1 ## s2" "one (1 + v) s2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4509
    by (auto elim:sep_conjE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4510
  from hh(1)[unfolded reps_rev[OF False]]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4511
  obtain s11 s12 s13 where hhh:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4512
    "(reps u (v - int (x + 1) - 1) xs) s11"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4513
    "(zero (v - int (x + 1))) s12" "(ones (v - int x) v) s13"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4514
    "s11 ## (s12 + s13)" "s12 ## s13" "s1 = s11 + s12 + s13"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4515
    apply (atomize_elim)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4516
    apply(elim sep_conjE)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4517
    apply (rule_tac x = xa in exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4518
    apply (rule_tac x = xaa in exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4519
    apply (rule_tac x = ya in exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4520
    apply (intro conjI, assumption+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4521
    by (auto simp:set_ins_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4522
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4523
  proof(rule reps_splitedI[where v = "(v - int (x + 1) - 1)"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4524
    show "(reps u (v - int (x + 1) - 1) xs \<and>* zero (v - int (x + 1) - 1 + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4525
                                    reps (v - int (x + 1) - 1 + 2) (1 + v) [Suc x]) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4526
    proof(rule sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4527
      from hhh(1) show "reps u (v - int (x + 1) - 1) xs s11" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4528
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4529
      show "(zero (v - int (x + 1) - 1 + 1) \<and>* reps (v - int (x + 1) - 1 + 2) (1 + v) [Suc x]) (s12 + (s13 + s2))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4530
      proof(rule sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4531
        from hhh(2) show "zero (v - int (x + 1) - 1 + 1) s12" by smt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4532
      next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4533
        from hh(4) hhh(3)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4534
        show "reps (v - int (x + 1) - 1 + 2) (1 + v) [Suc x] (s13 + s2)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4535
          apply (simp add:reps.simps del:ones_simps add:ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4536
          by (smt hh(3) hh(4) hhh(4) hhh(5) hhh(6) sep_add_disjD sep_conjI sep_disj_addI1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4537
      next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4538
        show "s12 ## s13 + s2" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4539
          by (metis hh(3) hhh(4) hhh(5) hhh(6) sep_add_commute sep_add_disjD 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4540
              sep_add_disjI2 sep_disj_addD sep_disj_addD1 sep_disj_addI1 sep_disj_commuteI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4541
      next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4542
        show "s12 + (s13 + s2) = s12 + (s13 + s2)" by metis 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4543
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4544
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4545
      show "s11 ## s12 + (s13 + s2)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4546
        by (metis hh(3) hhh(4) hhh(5) hhh(6) sep_disj_addD1 sep_disj_addI1 sep_disj_addI3)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4547
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4548
      show "s = s11 + (s12 + (s13 + s2))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4549
        by (smt hh(2) hh(3) hhh(4) hhh(5) hhh(6) sep_add_assoc sep_add_commute 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4550
             sep_add_disjD sep_add_disjI2 sep_disj_addD1 sep_disj_addD2 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4551
              sep_disj_addI1 sep_disj_addI3 sep_disj_commuteI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4552
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4553
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4554
    from False show "xs \<noteq> []" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4555
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4556
    show "[Suc x] \<noteq> []" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4557
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4558
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4559
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4560
lemma cond_expand: "(<cond> \<and>* p) s = (cond \<and> (p s))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4561
  by (metis (full_types) condD pasrt_def sep_conj_commuteI 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4562
             sep_conj_sep_emptyI sep_empty_def sep_globalise)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4563
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4564
lemma ones_rev1:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4565
  assumes "\<not> (1 + n) < m"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4566
  shows "(ones m n \<and>* one (1 + n)) = (ones m (1 + n))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4567
  by (insert ones_rev[OF assms, simplified], simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4568
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4569
lemma reps_one_abs:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4570
  assumes "(reps u v [k] \<and>* one w) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4571
          "w = v + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4572
  shows "(reps u w [Suc k]) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4573
  using assms unfolding assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4574
  apply (simp add:reps.simps sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4575
  apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4576
  apply (erule condE, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4577
  by (simp add:ones_rev sep_conj_ac, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4578
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4579
lemma reps'_reps_abs:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4580
  assumes "(reps' u v xs \<and>* reps w x ys) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4581
          "w = v + 1"  "ys \<noteq> []"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4582
  shows "(reps u x (xs@ys)) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4583
proof(cases "xs = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4584
  case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4585
  with assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4586
  have h: "splited (xs@ys) xs ys" by (simp add:splited_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4587
  from assms(1)[unfolded assms(2)]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4588
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4589
    apply (unfold reps_splited[OF h])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4590
    apply (insert False, unfold reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4591
    apply (fwd reps_lenE, (subst (asm) sep_conj_cond)+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4592
    by (erule condE, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4593
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4594
  case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4595
  with assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4596
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4597
    apply (simp add:reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4598
    by (erule condE, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4599
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4600
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4601
lemma reps_one_abs1:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4602
  assumes "(reps u v (xs@[k]) \<and>* one w) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4603
          "w = v + 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4604
  shows "(reps u w (xs@[Suc k])) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4605
proof(cases "xs = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4606
  case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4607
  with assms reps_one_abs
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4608
  show ?thesis by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4609
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4610
  case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4611
  hence "splited (xs@[k]) xs [k]" by (simp add:splited_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4612
  from assms(1)[unfolded reps_splited[OF this] assms(2)]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4613
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4614
    apply (fwd reps_one_abs)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4615
    apply (fwd reps_reps'_abs) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4616
    apply (fwd reps'_reps_abs)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4617
    by (simp add:assms)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4618
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4619
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4620
lemma tm_hoare_inc00: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4621
  assumes h: "a < length ks \<and> ks ! a = v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4622
  shows "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4623
    i :[ Inc a ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4624
    \<lbrace>st j \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4625
     ps u \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4626
     zero (u - 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4627
     zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4628
     reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) (ks[a := Suc v]) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4629
     fam_conj {(ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1)<..} zero\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4630
  (is "\<lbrace> ?P \<rbrace> ?code \<lbrace> ?Q \<rbrace>")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4631
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4632
  from h have "a < length ks" "ks ! a = v" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4633
  from list_nth_expand[OF `a < length ks`]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4634
  have eq_ks: "ks = take a ks @ [ks!a] @ drop (Suc a) ks" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4635
  from `a < length ks` have "ks \<noteq> []" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4636
  hence "(reps u ia ks \<and>* zero (ia + 1)) = reps' u (ia + 1) ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4637
    by (simp add:reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4638
  also from eq_ks have "\<dots> = reps' u (ia + 1) (take a ks @ [ks!a] @ drop (Suc a) ks)" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4639
  also have "\<dots>  = (EXS m. reps' u (m - 1) (take a ks) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4640
                     reps' m (ia + 1) (ks ! a # drop (Suc a) ks))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4641
    by (simp add:reps'_append)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4642
  also have "\<dots> = (EXS m. reps' u (m - 1) (take a ks) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4643
                     reps' m (ia + 1) ([ks ! a] @ drop (Suc a) ks))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4644
    by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4645
  also have "\<dots> = (EXS m ma. reps' u (m - 1) (take a ks) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4646
                       reps' m (ma - 1) [ks ! a] \<and>* reps' ma (ia + 1) (drop (Suc a) ks))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4647
    by (simp only:reps'_append sep_conj_exists)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4648
  finally have eq_s: "(reps u ia ks \<and>* zero (ia + 1)) = \<dots>" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4649
  { fix m ma
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4650
    have eq_u: "-1 + u = u - 1" by smt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4651
    have " \<lbrace>st i \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4652
            ps u \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4653
            zero (u - 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4654
            zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4655
            (reps' u (m - 1) (take a ks) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4656
             reps' m (ma - 1) [ks ! a] \<and>* reps' ma (ia + 1) (drop (Suc a) ks)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4657
            fam_conj {ia + 1<..} zero\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4658
           i :[ Inc a ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4659
           \<lbrace>st j \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4660
            ps u \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4661
            zero (u - 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4662
            zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4663
            reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) (ks[a := Suc v]) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4664
            fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1<..} zero\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4665
    proof(cases "(drop (Suc a) ks) = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4666
      case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4667
      { fix hc
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4668
        have eq_1: "(1 + (m + int (ks ! a))) = (m + int (ks ! a) + 1)" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4669
        have eq_2: "take a ks @ [Suc (ks ! a)] = ks[a := Suc v]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4670
          apply (subst (3) eq_ks, unfold True, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4671
          by (metis True append_Nil2 eq_ks h upd_conv_take_nth_drop)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4672
        assume h: "(fam_conj {1 + (m + int (ks ! a))<..} zero \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4673
                      reps u (1 + (m + int (ks ! a))) (take a ks @ [Suc (ks ! a)])) hc"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4674
        hence "(fam_conj {m + int (ks ! a) + 1<..} zero \<and>* reps u (m + int (ks ! a) + 1) (ks[a := Suc v])) hc"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4675
          by (unfold eq_1 eq_2 , sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4676
      } note eq_fam = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4677
      show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4678
        apply (unfold Inc_def, subst (3) reps'_def, simp add:True sep_conj_cond)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4679
        apply (intro tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4680
        apply (subst fam_conj_interv_simp, simp, subst (3) zero_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4681
        apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4682
        apply (subst reps'_sg, simp add:sep_conj_cond del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4683
        apply (rule tm.pre_condI, simp del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4684
        apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4685
        apply (rule_tac p = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4686
          st j' \<and>* ps (1 + (m + int (ks ! a))) \<and>* zero (u - 1) \<and>* zero (u - 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4687
                   reps u (1 + (m + int (ks ! a))) (take a ks @ [Suc (ks!a)]) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4688
                            \<and>* fam_conj {1 + (m + int (ks ! a))<..} zero
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4689
          " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4690
        apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4691
        apply (simp add:sep_conj_ac list_ext_lt[OF `a < length ks`], sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4692
        apply (fwd eq_fam, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4693
        apply (simp del:ones_simps add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4694
        apply (sep_cancel+, prune)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4695
        apply ((fwd abs_reps')+, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4696
        apply (fwd reps_one_abs abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4697
        apply (subst (asm) reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4698
        by (subst fam_conj_interv_simp, simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4699
    next 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4700
      case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4701
      then obtain k' ks' where eq_drop: "drop (Suc a) ks = [k']@ks'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4702
        by (metis append_Cons append_Nil list.exhaust)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4703
      from `a < length ks`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4704
      have eq_ks: "ks[a := Suc v] = (take a ks @ [Suc (ks ! a)]) @ (drop (Suc a) ks)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4705
        apply (fold `ks!a = v`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4706
        by (metis append_Cons append_Nil append_assoc upd_conv_take_nth_drop)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4707
      show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4708
        apply (unfold Inc_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4709
        apply (unfold Inc_def eq_drop reps'_append, simp add:sep_conj_exists del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4710
        apply (rule tm.precond_exI, subst (2) reps'_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4711
        apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4712
        apply (subst (1) ones_int_expand)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4713
        apply (rule tm.pre_condI, simp del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4714
        apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4715
        (* apply (hsteps hoare_locate_skip_gen[OF `a < length ks`]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4716
        apply (subst reps'_sg, simp add:sep_conj_cond del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4717
        apply (rule tm.pre_condI, simp del:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4718
        apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4719
        apply (rule_tac p = "st j' \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4720
                ps (2 + (m + int (ks ! a))) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4721
                reps (2 + (m + int (ks ! a))) ia (drop (Suc a) ks) \<and>* zero (ia + 1) \<and>* zero (ia + 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4722
                reps u (m + int (ks ! a)) (take a ks @ [ks!a]) \<and>* zero (1 + (m + int (ks ! a))) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4723
                zero (u - 2) \<and>* zero (u - 1) \<and>* fam_conj {ia + 2<..} zero
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4724
          " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4725
        apply (hsteps hoare_shift_right_cons_gen[OF False]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4726
                hoare_left_until_double_zero_gen[OF False])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4727
        apply (rule_tac p = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4728
          "st j' \<and>* ps (1 + (m + int (ks ! a))) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4729
          zero (u - 2) \<and>* zero (u - 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4730
          reps u (1 + (m + int (ks ! a))) (take a ks @ [Suc (ks ! a)]) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4731
          zero (2 + (m + int (ks ! a))) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4732
          reps (3 + (m + int (ks ! a))) (ia + 1) (drop (Suc a) ks) \<and>* fam_conj {ia + 1<..} zero
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4733
          " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4734
        apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4735
        apply (simp add:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4736
        apply (unfold list_ext_lt[OF `a < length ks`], simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4737
        apply (fwd abs_reps')+ 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4738
        apply(fwd reps'_reps_abs)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4739
        apply (subst eq_ks, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4740
        apply (sep_cancel+) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4741
        apply (thin_tac "mb = 4 + (m + (int (ks ! a) + int k'))")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4742
        apply (thin_tac "ma = 2 + (m + int (ks ! a))", prune)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4743
        apply (simp add: int_add_ac sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4744
        apply (fwd reps_one_abs1, subst fam_conj_interv_simp, simp, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4745
        apply (fwd abs_ones)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4746
        apply (fwd abs_reps')
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4747
        apply (fwd abs_reps')
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4748
        apply (fwd abs_reps')
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4749
        apply (fwd abs_reps')
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4750
        apply (unfold eq_drop, simp add:int_add_ac sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4751
        apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4752
        apply (fwd  reps'_abs[where xs = "take a ks"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4753
        apply (fwd reps'_abs[where xs = "[k']"])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4754
        apply (unfold reps'_def, simp add:int_add_ac sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4755
        apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4756
        by (subst (asm) fam_conj_interv_simp, simp, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4757
      qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4758
  } note h = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4759
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4760
    apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4761
    apply (rule_tac p = "st i \<and>*  ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4762
                              (reps u ia ks \<and>* zero (ia + 1)) \<and>* fam_conj {ia + 1<..} zero" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4763
      in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4764
    apply (unfold eq_s, simp only:sep_conj_exists)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4765
    apply (intro tm.precond_exI h)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4766
    by (sep_cancel+, unfold eq_s, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4767
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4768
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4769
declare ones_simps [simp del]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4770
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4771
lemma tm_hoare_inc01:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4772
  assumes "length ks \<le> a \<and> v = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4773
  shows 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4774
   "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4775
       i :[ Inc a ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4776
    \<lbrace>st j \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4777
     ps u \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4778
     zero (u - 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4779
     zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4780
     reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) ((list_ext a ks)[a := Suc v]) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4781
     fam_conj {(ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1)<..} zero\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4782
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4783
  from assms have "length ks \<le> a" "v = 0" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4784
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4785
    apply (rule_tac p = "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4786
      st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* (reps u ia ks \<and>* <(ia = u + int (reps_len ks) - 1)>) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4787
             fam_conj {ia<..} zero
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4788
      " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4789
    apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4790
    apply (rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4791
    apply (unfold Inc_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4792
    apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4793
    (* apply (hstep hoare_locate_set_gen[OF `length ks \<le> a`]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4794
    apply (simp only:sep_conj_exists)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4795
    apply (intro tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4796
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4797
      fix m w
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4798
      have "reps m w [list_ext a ks ! a] =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4799
            (ones m (m + int (list_ext a ks ! a)) \<and>* <(w = m + int (list_ext a ks ! a))>)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4800
        by (simp add:reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4801
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4802
    apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4803
    apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4804
    apply (rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4805
    apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4806
    apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4807
    apply (subst fam_conj_interv_simp, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4808
    apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4809
    apply (rule_tac p = "st j' \<and>* ps (m + int (list_ext a ks ! a) + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4810
                           zero (u - 2) \<and>* zero (u - 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4811
                           reps u (m + int (list_ext a ks ! a) + 1) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4812
                                ((take a (list_ext a ks))@[Suc (list_ext a ks ! a)]) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4813
                           fam_conj {(m + int (list_ext a ks ! a) + 1)<..} zero
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4814
                         " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4815
    apply (hsteps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4816
    apply (simp add:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4817
    apply (unfold `v = 0`, prune)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4818
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4819
      from `length ks \<le> a` have "list_ext a ks ! a = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4820
        by (metis le_refl list_ext_tail)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4821
      from `length ks \<le> a` have "a < length (list_ext a ks)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4822
        by (metis list_ext_len)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4823
      from reps_len_suc[OF this] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4824
      have eq_len: "int (reps_len (list_ext a ks)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4825
                        int (reps_len (list_ext a ks[a := Suc (list_ext a ks ! a)])) - 1" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4826
        by smt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4827
      fix m w hc
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4828
      assume h: "(fam_conj {m + int (list_ext a ks ! a) + 1<..} zero \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4829
                 reps u (m + int (list_ext a ks ! a) + 1) (take a (list_ext a ks) @ [Suc (list_ext a ks ! a)]))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4830
                 hc"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4831
      then obtain s where 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4832
        "(reps u (m + int (list_ext a ks ! a) + 1) (take a (list_ext a ks) @ [Suc (list_ext a ks ! a)])) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4833
        by (auto dest!:sep_conjD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4834
      from reps_len_correct[OF this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4835
      have "m  = u + int (reps_len (take a (list_ext a ks) @ [Suc (list_ext a ks ! a)])) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4836
                        - int (list_ext a ks ! a) - 2" by smt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4837
      from h [unfolded this]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4838
      have "(fam_conj {u + int (reps_len (list_ext a ks))<..} zero \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4839
           reps u (u + int (reps_len (list_ext a ks))) (list_ext a ks[a := Suc 0]))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4840
           hc"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4841
        apply (unfold eq_len, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4842
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4843
          from `a < length (list_ext a ks)`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4844
          have "take a (list_ext a ks) @ [Suc (list_ext a ks ! a)] = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4845
                list_ext a ks[a := Suc (list_ext a ks ! a)]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4846
            by (smt `list_ext a ks ! a = 0` assms length_take list_ext_tail_expand list_update_length)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4847
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4848
        apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4849
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4850
          have "-1 + (u + int (reps_len (list_ext a ks[a := Suc (list_ext a ks ! a)]))) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4851
                u + (int (reps_len (list_ext a ks[a := Suc (list_ext a ks ! a)])) - 1)" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4852
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4853
        apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4854
        apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4855
        by (unfold `(list_ext a ks ! a) = 0`, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4856
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4857
    apply (rule this, assumption)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4858
    apply (simp only:sep_conj_ac, sep_cancel+)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4859
    apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4860
    apply (fwd reps_one_abs) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4861
    apply (fwd reps'_reps_abs)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4862
    apply (simp add:int_add_ac sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4863
    apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4864
    apply (subst fam_conj_interv_simp, simp add:sep_conj_ac, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4865
    apply (fwd reps_lenE, (subst (asm) sep_conj_cond)+, erule condE, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4866
    by (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4867
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4868
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4869
definition "Dec a e  = (TL continue. 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4870
                          (locate a; 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4871
                           if_reps_nz continue;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4872
                           left_until_double_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4873
                           move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4874
                           move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4875
                           jmp e);
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4876
                          (TLabel continue;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4877
                           right_until_zero; 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4878
                           move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4879
                           write_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4880
                           move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4881
                           move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4882
                           shift_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4883
                           move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4884
                           move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4885
                           move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4886
                           left_until_double_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4887
                           move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4888
                           move_right))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4889
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4890
lemma cond_eq: "((<b> \<and>* p) s) = (b \<and> (p s))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4891
proof
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4892
  assume "(<b> \<and>* p) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4893
  from condD[OF this] show " b \<and> p s" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4894
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4895
  assume "b \<and> p s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4896
  hence b and "p s" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4897
  from `b` have "(<b>) 0" by (auto simp:pasrt_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4898
  moreover have "s = 0 + s" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4899
  moreover have "0 ## s" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4900
  moreover note `p s`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4901
  ultimately show "(<b> \<and>* p) s" by (auto intro!:sep_conjI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4902
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4903
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4904
lemma tm_hoare_dec_fail00:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4905
  assumes "a < length ks \<and> ks ! a = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4906
  shows "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace>  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4907
             i :[ Dec a e ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4908
         \<lbrace>st e \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4909
          reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4910
          fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4911
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4912
  from assms have "a < length ks" "ks!a = 0" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4913
  from list_nth_expand[OF `a < length ks`]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4914
  have eq_ks: "ks = take a ks @ [ks ! a] @ drop (Suc a) ks" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4915
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4916
  proof(cases " drop (Suc a) ks = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4917
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4918
    then obtain k' ks' where eq_drop: "drop (Suc a) ks = [k']@ks'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4919
      by (metis append_Cons append_Nil list.exhaust)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4920
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4921
      apply (unfold Dec_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4922
      apply (subst tassemble_to.simps(2), rule tm.code_exI, rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4923
      apply (subst (1) eq_ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4924
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4925
        have "(reps u ia (take a ks @ [ks ! a] @ drop (Suc a) ks) \<and>* fam_conj {ia<..} zero) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4926
              (reps' u (ia + 1) ((take a ks @ [ks ! a]) @ drop (Suc a) ks) \<and>* fam_conj {ia + 1<..} zero)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4927
          apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4928
          by (unfold reps'_def, simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4929
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4930
      apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4931
      apply (subst reps'_append)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4932
      apply (unfold eq_drop)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4933
      apply (subst (2) reps'_append)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4934
      apply (simp only:sep_conj_exists, intro tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4935
      apply (subst (2) reps'_def, simp add:reps.simps ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4936
      apply (subst reps'_append, simp only:sep_conj_exists, intro tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4937
      apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4938
      (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4939
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4940
        fix m mb
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4941
        have "(reps' mb (m - 1) [ks ! a]) = (reps mb (m - 2) [ks!a] \<and>* zero (m - 1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4942
          by (simp add:reps'_def, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4943
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4944
      apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4945
      apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4946
      (* apply (hstep hoare_if_reps_nz_false_gen[OF `ks!a = 0`]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4947
      apply (simp only:reps.simps(2), simp add:`ks!a = 0`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4948
      apply (rule_tac p = "st j'b \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4949
        ps mb \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4950
        reps u mb ((take a ks)@[ks ! a]) \<and>* <(m - 2 = mb)> \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4951
        zero (m - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4952
        zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4953
        one m \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4954
        zero (u - 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4955
        ones (m + 1) (m + int k') \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4956
        <(-2 + ma = m + int k')> \<and>* zero (ma - 1) \<and>* reps' ma (ia + 1) ks' \<and>* fam_conj {ia + 1<..} zero"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4957
        in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4958
      apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4959
      apply (simp add:sep_conj_ac, sep_cancel+) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4960
      apply (subgoal_tac "m + int k' = ma - 2", simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4961
      apply (fwd abs_ones)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4962
      apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4963
      apply (erule condE, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4964
      apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4965
      apply (subgoal_tac "ma = m + int k' + 2", simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4966
      apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4967
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4968
        from `a < length ks`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4969
        have "list_ext a ks = ks" by (auto simp:list_ext_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4970
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4971
      apply (simp add:this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4972
      apply (subst eq_ks, simp add:eq_drop `ks!a = 0`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4973
      apply (subst (asm) reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4974
      apply (subst fam_conj_interv_simp, simp add:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4975
      apply (metis append_Cons assms eq_Nil_appendI eq_drop eq_ks list_update_id)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4976
      apply (clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4977
      apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4978
      apply (erule condE, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4979
      apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4980
      apply (erule condE, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4981
      apply (simp add:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4982
      apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4983
      by (fwd reps'_reps_abs, simp add:`ks!a = 0`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4984
  next 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4985
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4986
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4987
      apply (unfold Dec_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4988
      apply (subst tassemble_to.simps(2), rule tm.code_exI, rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4989
      apply (subst (1) eq_ks, unfold True, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4990
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4991
        have "(reps u ia (take a ks @ [ks ! a]) \<and>* fam_conj {ia<..} zero) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4992
              (reps' u (ia + 1) (take a ks @ [ks ! a]) \<and>* fam_conj {ia + 1<..} zero)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4993
          apply (unfold reps'_def, subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4994
          by (simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4995
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4996
      apply (subst (1) this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4997
      apply (subst reps'_append)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4998
      apply (simp only:sep_conj_exists, intro tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  4999
      apply (subst fam_conj_interv_simp, simp) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5000
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5001
        have "(zero (2 + ia)) = (tm (2 + ia) Bk)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5002
          by (simp add:zero_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5003
      my_block_end my_note eq_z = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5004
      apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5005
      (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5006
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5007
        fix m 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5008
        have "(reps' m (ia + 1) [ks ! a]) = (reps m ia [ks!a] \<and>* zero (ia + 1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5009
          by (simp add:reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5010
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5011
      apply (unfold this, prune)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5012
      apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5013
      (* apply (hstep hoare_if_reps_nz_false_gen[OF `ks!a = 0`]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5014
      apply (simp only:reps.simps(2), simp add:`ks!a = 0`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5015
      apply (rule_tac p = "st j'b \<and>* ps m \<and>* (reps u m ((take a ks)@[ks!a]) \<and>* <(ia = m)>) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5016
                              \<and>* zero (ia + 1) \<and>* zero (u - 1) \<and>*  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5017
                              zero (2 + ia) \<and>* zero (u - 2) \<and>* fam_conj {2 + ia<..} zero"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5018
        in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5019
      apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5020
      apply (simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5021
      apply ((subst (asm) sep_conj_cond)+, erule condE, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5022
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5023
        from `a < length ks`  have "list_ext a ks = ks" by (metis list_ext_lt) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5024
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5025
      apply (unfold this, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5026
      apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5027
      apply (subst fam_conj_interv_simp, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5028
      apply (simp only:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5029
      apply (subst eq_ks, unfold True `ks!a = 0`, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5030
      apply (metis True append_Nil2 assms eq_ks list_update_same_conv) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5031
      apply (simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5032
      apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5033
      apply (erule condE, simp, thin_tac "ia = m")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5034
      apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5035
      apply (simp add:sep_conj_ac int_add_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5036
      apply (unfold reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5037
      by (metis sep.mult_commute)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5038
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5039
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5040
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5041
lemma tm_hoare_dec_fail01:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5042
  assumes "length ks \<le> a"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5043
  shows "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace>  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5044
                       i :[ Dec a e ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5045
         \<lbrace>st e \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5046
          reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5047
          fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5048
  apply (unfold Dec_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5049
  apply (subst tassemble_to.simps(2), rule tm.code_exI, rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5050
  apply (rule_tac p = "st i \<and>* ps u \<and>* zero (u - 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5051
                       zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5052
                       <(ia = u + int (reps_len ks) - 1)>" in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5053
  apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5054
  (* apply (hstep hoare_locate_set_gen[OF `length ks \<le> a`]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5055
  apply (simp only:sep_conj_exists, intro tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5056
  my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5057
    from assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5058
    have "list_ext a ks ! a = 0" by (metis le_refl list_ext_tail) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5059
  my_block_end my_note is_z = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5060
  apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5061
  apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5062
  (* apply (hstep hoare_if_reps_nz_false_gen[OF is_z]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5063
  apply (unfold is_z)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5064
  apply (subst (1) reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5065
  apply (rule_tac p = "st j'b \<and>* ps m \<and>*  reps u m (take a (list_ext a ks) @ [0]) \<and>* zero (w + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5066
                         <(w = m + int 0)> \<and>* zero (u - 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5067
                         fam_conj {w + 1<..} zero \<and>* zero (u - 2) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5068
                         <(ia = u + int (reps_len ks) - 1)>" in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5069
  my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5070
    have "(take a (list_ext a ks)) @ [0] \<noteq> []" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5071
  my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5072
  apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5073
  (* apply (hsteps hoare_left_until_double_zero_gen[OF this]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5074
  apply (simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5075
  apply prune
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5076
  apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5077
  apply (elim condE, simp add:sep_conj_ac, prune)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5078
  my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5079
    fix m w ha
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5080
    assume h1: "w = m \<and> ia = u + int (reps_len ks) - 1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5081
      and  h: "(ps u \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5082
              st e \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5083
              zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5084
              zero (m + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5085
              fam_conj {m + 1<..} zero \<and>* zero (u - 2) \<and>* reps u m (take a (list_ext a ks) @ [0])) ha"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5086
    from h1 have eq_w: "w = m" and eq_ia: "ia = u + int (reps_len ks) - 1" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5087
    from h obtain s' where "reps u m (take a (list_ext a ks) @ [0]) s'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5088
      by (auto dest!:sep_conjD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5089
    from reps_len_correct[OF this] 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5090
    have eq_m: "m = u + int (reps_len (take a (list_ext a ks) @ [0])) - 1" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5091
    from h[unfolded eq_m, simplified]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5092
    have "(ps u \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5093
                st e \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5094
                zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5095
                zero (u - 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5096
                fam_conj {u + (-1 + int (reps_len (list_ext a ks)))<..} zero \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5097
                reps u (u + (-1 + int (reps_len (list_ext a ks)))) (list_ext a ks[a := 0])) ha"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5098
      apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5099
      apply (subst fam_conj_interv_simp, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5100
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5101
        from `length ks \<le> a` have "list_ext a ks[a := 0] = list_ext a ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5102
          by (metis is_z list_update_id)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5103
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5104
      apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5105
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5106
        from `length ks \<le> a` is_z 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5107
        have "take a (list_ext a ks) @ [0] = list_ext a ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5108
          by (metis list_ext_tail_expand)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5109
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5110
      apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5111
      by (simp add:sep_conj_ac, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5112
  my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5113
  apply (rule this, assumption)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5114
  apply (sep_cancel+)[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5115
  apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5116
  apply (erule condE, prune, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5117
  my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5118
    fix s m
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5119
    assume "(reps' u (m - 1) (take a (list_ext a ks)) \<and>* ones m m \<and>* zero (m + 1)) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5120
    hence "reps' u (m + 1) (take a (list_ext a ks) @ [0]) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5121
      apply (unfold reps'_append)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5122
      apply (rule_tac x = m in EXS_intro)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5123
      by (subst (2) reps'_def, simp add:reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5124
  my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5125
  apply (rotate_tac 1, fwd this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5126
  apply (subst (asm) reps'_def, simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5127
  my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5128
    fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5129
    assume h: "(st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5130
                   reps u ia ks \<and>* fam_conj {ia<..} zero) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5131
    then obtain s' where "reps u ia ks s'" by (auto dest!:sep_conjD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5132
    from reps_len_correct[OF this] have eq_ia: "ia = u + int (reps_len ks) - 1" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5133
    from h
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5134
    have "(st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5135
           fam_conj {ia<..} zero \<and>* <(ia = u + int (reps_len ks) - 1)>) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5136
      by (unfold eq_ia, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5137
  my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5138
  by (rule this, assumption)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5139
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5140
lemma t_hoare_label1: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5141
      "(\<And>l. l = i \<Longrightarrow> \<lbrace>st l \<and>* p\<rbrace>  l :[ c l ]: j \<lbrace>st k \<and>* q\<rbrace>) \<Longrightarrow>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5142
             \<lbrace>st l \<and>* p \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5143
               i:[(TLabel l; c l)]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5144
             \<lbrace>st k \<and>* q\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5145
by (unfold tassemble_to.simps, intro tm.code_exI tm.code_condI, clarify, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5146
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5147
lemma tm_hoare_dec_fail1:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5148
  assumes "a < length ks \<and> ks ! a = 0 \<or> length ks \<le> a"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5149
  shows "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace>  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5150
                      i :[ Dec a e ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5151
         \<lbrace>st e \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5152
          reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5153
         fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5154
  using assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5155
proof
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5156
  assume "a < length ks \<and> ks ! a = 0"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5157
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5158
    by (rule tm_hoare_dec_fail00)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5159
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5160
  assume "length ks \<le> a"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5161
  thus ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5162
    by (rule tm_hoare_dec_fail01)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5163
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5164
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5165
lemma shift_left_nil_gen[step]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5166
  assumes "u = v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5167
  shows "\<lbrace>st i \<and>* ps u \<and>* zero v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5168
              i :[shift_left]:j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5169
         \<lbrace>st j \<and>* ps u \<and>* zero v\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5170
 apply(unfold assms shift_left_def, intro t_hoare_local t_hoare_label, clarify, 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5171
                 rule t_hoare_label_last, simp, clarify, prune, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5172
 by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5173
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5174
lemma tm_hoare_dec_suc1: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5175
  assumes "a < length ks \<and> ks ! a = Suc v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5176
  shows "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5177
                    i :[ Dec a e ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5178
         \<lbrace>st j \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5179
             reps u (ia - 1) (list_ext a ks[a := v]) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5180
             fam_conj {ia - 1<..} zero\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5181
proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5182
  from assms have "a < length ks" " ks ! a = Suc v" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5183
  from list_nth_expand[OF `a < length ks`]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5184
  have eq_ks: "ks = take a ks @ [ks ! a] @ drop (Suc a) ks" .
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5185
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5186
  proof(cases " drop (Suc a) ks = []")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5187
    case False
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5188
    then obtain k' ks' where eq_drop: "drop (Suc a) ks = [k']@ks'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5189
      by (metis append_Cons append_Nil list.exhaust)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5190
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5191
      apply (unfold Dec_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5192
      apply (subst tassemble_to.simps(2), rule tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5193
      apply (subst (1) eq_ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5194
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5195
        have "(reps u ia (take a ks @ [ks ! a] @ drop (Suc a) ks) \<and>* fam_conj {ia<..} zero) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5196
              (reps' u (ia + 1) ((take a ks @ [ks ! a]) @ drop (Suc a) ks) \<and>* fam_conj {ia + 1<..} zero)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5197
          apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5198
          by (unfold reps'_def, simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5199
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5200
      apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5201
      apply (subst reps'_append)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5202
      apply (unfold eq_drop)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5203
      apply (subst (2) reps'_append)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5204
      apply (simp only:sep_conj_exists, intro tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5205
      apply (subst (2) reps'_def, simp add:reps.simps ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5206
      apply (subst reps'_append, simp only:sep_conj_exists, intro tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5207
      apply (rule_tac q =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5208
       "st l \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5209
        ps mb \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5210
        zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5211
        reps' u (mb - 1) (take a ks) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5212
        reps' mb (m - 1) [ks ! a] \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5213
        one m \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5214
        zero (u - 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5215
        ones (m + 1) (m + int k') \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5216
        <(-2 + ma = m + int k')> \<and>* zero (ma - 1) \<and>* reps' ma (ia + 1) ks' \<and>* fam_conj {ia + 1<..} zero"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5217
      in tm.sequencing)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5218
      apply (rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5219
      apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5220
      (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5221
      apply (subst (2) reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5222
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5223
        fix i j l m mb
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5224
        from `ks!a = (Suc v)` have "ks!a \<noteq> 0" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5225
        from hoare_if_reps_nz_true[OF this, where u = mb and v = "m - 2"]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5226
        have "\<lbrace>st i \<and>* ps mb \<and>* reps mb (-2 + m) [ks ! a]\<rbrace>  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5227
                        i :[ if_reps_nz l ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5228
              \<lbrace>st l \<and>* ps mb \<and>* reps mb (-2 + m) [ks ! a]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5229
          by smt
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5230
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5231
      apply (hgoto this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5232
      apply (simp add:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5233
      apply (subst reps'_def, simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5234
      apply (rule tm.code_extension1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5235
      apply (rule t_hoare_label1, simp, prune)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5236
      apply (subst (2) reps'_def, simp add:reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5237
      apply (rule_tac p = "st j' \<and>* ps mb \<and>* zero (u - 1) \<and>* reps' u (mb - 1) (take a ks) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5238
        ((ones mb (mb + int (ks ! a)) \<and>* <(-2 + m = mb + int (ks ! a))>) \<and>* zero (mb + int (ks ! a) + 1)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5239
          one (mb + int (ks ! a) + 2) \<and>* zero (u - 2) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5240
          ones (mb + int (ks ! a) + 3) (mb + int (ks ! a) + int k' + 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5241
        <(-2 + ma = m + int k')> \<and>* zero (ma - 1) \<and>* reps' ma (ia + 1) ks' \<and>* fam_conj {ia + 1<..} zero
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5242
        " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5243
      apply hsteps 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5244
      (* apply (simp add:sep_conj_ac) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5245
      apply (unfold `ks!a = Suc v`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5246
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5247
        fix mb
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5248
        have "(ones mb (mb + int (Suc v))) = (ones mb (mb + int v) \<and>* one (mb + int (Suc v)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5249
          by (simp add:ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5250
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5251
      apply (unfold this, prune)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5252
      apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5253
      apply (rule_tac p = "st j'a \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5254
               ps (mb + int (Suc v) + 2) \<and>* zero (mb + int (Suc v) + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5255
               reps (mb + int (Suc v) + 2) ia (drop (Suc a) ks) \<and>* zero (ia + 1) \<and>* zero (ia + 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5256
        zero (mb + int (Suc v)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5257
        ones mb (mb + int v) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5258
        zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5259
        reps' u (mb - 1) (take a ks) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5260
        zero (u - 2) \<and>* fam_conj {ia + 2<..} zero
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5261
        " in tm.pre_stren) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5262
      apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5263
      (* apply (hsteps hoare_shift_left_cons_gen[OF False]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5264
      apply (rule_tac p = "st j'a \<and>* ps (ia - 1) \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5265
                           reps u (ia - 1) (take a ks @ [v] @ drop (Suc a) ks) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5266
                           zero ia \<and>* zero (ia + 1) \<and>* zero (ia + 2) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5267
                           fam_conj {ia + 2<..} zero
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5268
        " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5269
      apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5270
      apply (simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5271
      apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5272
      apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5273
      apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5274
      apply (simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5275
      apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5276
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5277
        have "take a ks @ v # drop (Suc a) ks = list_ext a ks[a := v]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5278
        proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5279
          from `a < length ks` have "list_ext a ks = ks" by (metis list_ext_lt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5280
          hence "list_ext a ks[a:=v] = ks[a:=v]" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5281
          moreover from `a < length ks` have "ks[a:=v] = take a ks @ v # drop (Suc a) ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5282
            by (metis upd_conv_take_nth_drop)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5283
          ultimately show ?thesis by metis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5284
        qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5285
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5286
      apply (unfold this, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5287
      apply (simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5288
      apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5289
      apply (simp add:sep_conj_ac int_add_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5290
      apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5291
      apply (subst (asm) reps'_def, simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5292
      apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5293
      apply (erule condE, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5294
      apply (simp add:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5295
      apply (fwd abs_ones)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5296
      apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5297
      apply (subst (asm) reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5298
      apply (subst (asm) fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5299
      apply (simp add:sep_conj_ac int_add_ac eq_drop reps'_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5300
      apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5301
      apply (erule condE, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5302
      by (simp add:sep_conj_ac int_add_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5303
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5304
    case True
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5305
    show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5306
      apply (unfold Dec_def, intro t_hoare_local)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5307
      apply (subst tassemble_to.simps(2), rule tm.code_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5308
      apply (subst (1) eq_ks, simp add:True)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5309
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5310
        have "(reps u ia (take a ks @ [ks ! a]) \<and>* fam_conj {ia<..} zero) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5311
              (reps' u (ia + 1) (take a ks @ [ks ! a]) \<and>* fam_conj {ia + 1<..} zero)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5312
          apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5313
          by (unfold reps'_def, simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5314
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5315
      apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5316
      apply (subst reps'_append)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5317
      apply (simp only:sep_conj_exists, intro tm.precond_exI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5318
      apply (rule_tac q = "st l \<and>* ps m \<and>* zero (u - 1) \<and>* reps' u (m - 1) (take a ks) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5319
            reps' m (ia + 1) [ks ! a] \<and>* zero (2 + ia) \<and>* zero (u - 2) \<and>* fam_conj {2 + ia<..} zero"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5320
             in tm.sequencing)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5321
      apply (rule tm.code_extension)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5322
      apply (subst fam_conj_interv_simp, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5323
      apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5324
      (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5325
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5326
        fix m
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5327
        have "(reps' m (ia + 1) [ks ! a]) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5328
              (reps m ia [ks!a] \<and>* zero (ia + 1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5329
          by (unfold reps'_def, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5330
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5331
      apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5332
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5333
        fix i j l m
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5334
        from `ks!a = (Suc v)` have "ks!a \<noteq> 0" by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5335
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5336
      apply (hgoto hoare_if_reps_nz_true_gen)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5337
      apply (rule tm.code_extension1)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5338
      apply (rule t_hoare_label1, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5339
      apply (thin_tac "la = j'", prune)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5340
      apply (subst (1) reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5341
      apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5342
      apply (rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5343
      apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5344
      apply (unfold `ks!a = Suc v`)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5345
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5346
        fix m
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5347
        have "(ones m (m + int (Suc v))) = (ones m (m + int v) \<and>* one (m + int (Suc v)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5348
          by (simp add:ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5349
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5350
      apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5351
      apply hsteps 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5352
      apply (rule_tac p = "st j'a \<and>* ps (m + int v) \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5353
                           reps u (m + int v) (take a ks @ [v]) \<and>* zero (m + (1 + int v)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5354
                           zero (2 + (m + int v)) \<and>* zero (3 + (m + int v)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5355
                           fam_conj {3 + (m + int v)<..} zero
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5356
        " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5357
      apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5358
      apply (simp add:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5359
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5360
        have "take a ks @ [v] = list_ext a ks[a := v]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5361
        proof -
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5362
          from True `a < length ks` have "ks = take a ks @ [ks!a]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5363
            by (metis append_Nil2 eq_ks)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5364
          hence "ks[a:=v] = take a ks @ [v]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5365
            by (metis True `a < length ks` upd_conv_take_nth_drop)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5366
          moreover from `a < length ks` have "list_ext a ks = ks"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5367
            by (metis list_ext_lt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5368
          ultimately show ?thesis by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5369
        qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5370
      my_block_end my_note eq_l = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5371
      apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5372
      apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5373
      apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5374
      apply (subst fam_conj_interv_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5375
      apply (simp add:sep_conj_ac, sep_cancel, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5376
      apply (simp add:sep_conj_ac int_add_ac)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5377
      apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5378
      apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5379
      apply (fwd reps'_reps_abs)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5380
      by (simp add:eq_l)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5381
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5382
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5383
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5384
definition "cfill_until_one = (TL start exit.
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5385
                                TLabel start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5386
                                  if_one exit;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5387
                                  write_one;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5388
                                  move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5389
                                  jmp start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5390
                                TLabel exit
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5391
                               )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5392
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5393
lemma hoare_cfill_until_one:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5394
   "\<lbrace>st i \<and>* ps v \<and>* one (u - 1) \<and>* zeros u v\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5395
              i :[ cfill_until_one ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5396
    \<lbrace>st j \<and>* ps (u - 1) \<and>* ones (u - 1) v \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5397
proof(induct u v rule:zeros_rev_induct)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5398
  case (Base x y)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5399
  thus ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5400
    apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5401
    apply (rule tm.pre_condI, simp add:ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5402
    apply (unfold cfill_until_one_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5403
    by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5404
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5405
  case (Step x y)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5406
  show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5407
    apply (rule_tac q = "st i \<and>* ps (y - 1) \<and>* one (x - 1) \<and>* zeros x (y - 1) \<and>* one y" in tm.sequencing)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5408
    apply (subst cfill_until_one_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5409
    apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5410
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5411
      fix i j l
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5412
      have "\<lbrace>st i \<and>* ps (y - 1) \<and>* one (x - 1) \<and>* zeros x (y - 1)\<rbrace>  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5413
              i :[ jmp l ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5414
            \<lbrace>st l \<and>* ps (y - 1) \<and>* one (x - 1) \<and>* zeros x (y - 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5415
        apply (case_tac "(y - 1) < x", simp add:zeros_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5416
        apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5417
        apply (rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5418
        apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5419
        apply (drule_tac zeros_rev, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5420
        by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5421
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5422
    apply (hstep this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5423
    (* The next half *)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5424
    apply (hstep Step(2), simp add:sep_conj_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5425
    by (insert Step(1), simp add:ones_rev sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5426
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5427
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5428
definition "cmove = (TL start exit.
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5429
                       TLabel start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5430
                         left_until_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5431
                         left_until_one;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5432
                         move_left;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5433
                         if_zero exit;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5434
                         move_right;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5435
                         write_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5436
                         right_until_one;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5437
                         right_until_zero;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5438
                         write_one;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5439
                         jmp start;
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5440
                     TLabel exit
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5441
                    )"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5442
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5443
declare zeros.simps [simp del] zeros_simps[simp del]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5444
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5445
lemma hoare_cmove:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5446
  assumes "w \<le> k"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5447
  shows "\<lbrace>st i \<and>* ps (v + 2 + int w) \<and>* zero (u - 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5448
              reps u (v - int w) [k - w] \<and>* zeros (v - int w + 1) (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5449
              one (v + 2) \<and>* ones (v + 3) (v + 2 + int w) \<and>* zeros (v + 3 + int w)  (v + int(reps_len [k]) + 1)\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5450
                                 i :[cmove]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5451
          \<lbrace>st j \<and>* ps (u - 1) \<and>* zero (u - 1) \<and>* reps u u [0] \<and>* zeros (u + 1) (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5452
                                                                  reps (v + 2) (v + int (reps_len [k]) + 1) [k]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5453
  using assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5454
proof(induct "k - w" arbitrary: w)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5455
  case (0 w)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5456
  hence "w = k" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5457
  show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5458
    apply (simp add: `w = k` del:zeros.simps zeros_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5459
    apply (unfold cmove_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5460
    apply (simp add:reps_len_def reps_sep_len_def reps_ctnt_len_def del:zeros_simps zeros.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5461
    apply (rule_tac p = "st i \<and>* ps (v + 2 + int k) \<and>* zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5462
                         reps u u [0] \<and>* zeros (u + 1) (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5463
                         ones (v + 2) (v + 2 + int k) \<and>* zeros (v + 3 + int k) (2 + (v + int k)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5464
                         <(u = v - int k)>" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5465
      in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5466
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5467
      fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5468
      have "\<lbrace>st i \<and>* ps (v + 2 + int k) \<and>* zeros (u + 1) (v + 1) \<and>* ones (v + 2) (v + 2 + int k) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5469
                                                             \<and>* <(u = v - int k)>\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5470
                  i :[ left_until_zero ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5471
            \<lbrace>st j \<and>* ps (v + 1) \<and>* zeros (u + 1) (v + 1) \<and>* ones (v + 2) (v + 2 + int k)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5472
                                                             \<and>* <(u = v - int k)>\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5473
        apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5474
        apply (rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5475
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5476
          have "(zeros (v - int k + 1) (v + 1)) = (zeros (v - int k + 1) v \<and>* zero (v + 1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5477
            by (simp only:zeros_rev, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5478
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5479
        apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5480
        by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5481
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5482
    apply (hstep this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5483
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5484
      fix i j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5485
      have "\<lbrace>st i \<and>* ps (v + 1) \<and>* reps u u [0] \<and>* zeros (u + 1) (v + 1)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5486
                i :[left_until_one]:j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5487
            \<lbrace>st j \<and>* ps u \<and>* reps u u [0] \<and>* zeros (u + 1) (v + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5488
        apply (simp add:reps.simps ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5489
        by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5490
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5491
    apply (hsteps this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5492
    apply ((subst (asm) sep_conj_cond)+, erule condE, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5493
    apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5494
    apply (simp only:sep_conj_ac int_add_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5495
    apply (simp add:int_add_ac sep_conj_ac zeros_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5496
    apply (simp add:sep_conj_ac int_add_ac, sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5497
    apply (fwd reps_lenE)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5498
    apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5499
    apply (erule condE, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5500
    apply (subgoal_tac "v  = u + int k + int (reps_len [0]) - 1", clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5501
    apply (simp add:reps_len_sg)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5502
    apply (fwd abs_ones)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5503
    apply (fwd abs_reps')+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5504
    apply (simp add:sep_conj_ac int_add_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5505
    apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5506
    apply (simp add:reps.simps, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5507
    by (clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5508
next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5509
  case (Suc k' w)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5510
  from `Suc k' = k - w` `w \<le> k` 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5511
  have h: "k' = k - (Suc w)" "Suc w \<le> k" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5512
  show ?case
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5513
    apply (rule tm.sequencing[OF _ Suc(1)[OF h(1, 2)]])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5514
    apply (unfold cmove_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5515
    apply (simp add:reps_len_def reps_sep_len_def reps_ctnt_len_def del:zeros_simps zeros.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5516
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5517
      fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5518
      have "\<lbrace>st i \<and>* ps (v + 2 + int w) \<and>* zeros (v - int w + 1) (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5519
                               one (v + 2) \<and>* ones (v + 3) (v + 2 + int w) \<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5520
                    i :[left_until_zero]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5521
            \<lbrace>st j \<and>* ps (v + 1) \<and>* zeros (v - int w + 1) (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5522
                               one (v + 2) \<and>* ones (v + 3) (v + 2 + int w) \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5523
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5524
          have "(one (v + 2) \<and>* ones (v + 3) (v + 2 + int w)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5525
                 ones (v + 2) (v + 2 + int w)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5526
            by (simp only:ones_simps, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5527
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5528
        apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5529
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5530
          have "(zeros (v - int w + 1) (v + 1)) = (zeros (v - int w + 1) v \<and>*  zero (v + 1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5531
            by (simp only:zeros_rev, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5532
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5533
        apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5534
        by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5535
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5536
    apply (hstep this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5537
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5538
      fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5539
      have "\<lbrace>st i \<and>* ps (v + 1) \<and>* reps u (v - int w) [k - w] \<and>* zeros (v - int w + 1) (v + 1)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5540
                 i :[left_until_one]: j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5541
            \<lbrace>st j \<and>* ps (v - int w) \<and>* reps u (v - int w) [k - w] \<and>* zeros (v - int w + 1) (v + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5542
        apply (simp add:reps.simps ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5543
        apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5544
        apply (rule tm.pre_condI, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5545
        apply (subgoal_tac "u + int (k - w) = v - int w", simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5546
        defer
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5547
        apply simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5548
        by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5549
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5550
    apply (hstep this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5551
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5552
      have "(reps u (v - int w) [k - w]) = (reps u (v - (1 + int w)) [k - Suc w] \<and>* one (v - int w))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5553
        apply (subst (1 2) reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5554
        apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5555
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5556
          have "((v - int w = u + int (k - w))) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5557
                (v - (1 + int w) = u + int (k - Suc w))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5558
            apply auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5559
            apply (smt Suc.prems h(2))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5560
            by (smt Suc.prems h(2))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5561
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5562
        apply (simp add:this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5563
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5564
          fix b p q
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5565
          assume "(b \<Longrightarrow> (p::tassert) = q)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5566
          have "(<b> \<and>* p) = (<b> \<and>* q)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5567
            by (metis `b \<Longrightarrow> p = q` cond_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5568
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5569
        apply (rule this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5570
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5571
          assume "v - (1 + int w) = u + int (k - Suc w)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5572
          hence "v = 1 + int w +  u + int (k - Suc w)" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5573
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5574
        apply (simp add:this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5575
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5576
          have "\<not> (u + int (k - w)) < u" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5577
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5578
        apply (unfold ones_rev[OF this])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5579
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5580
          from Suc (2, 3) have "(u + int (k - w) - 1) = (u + int (k - Suc w))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5581
            by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5582
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5583
        apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5584
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5585
          from Suc (2, 3) have "(u + int (k - w)) =  (1 + (u + int (k - Suc w)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5586
            by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5587
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5588
        by (unfold this, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5589
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5590
    apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5591
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5592
      fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5593
      have "\<lbrace>st i \<and>* ps (v - int w) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5594
                        (reps u (v - (1 + int w)) [k - Suc w] \<and>* one (v - int w))\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5595
                 i :[ move_left]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5596
            \<lbrace>st j \<and>* ps (v - (1 + int w)) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5597
                        (reps u (v - (1 + int w)) [k - Suc w] \<and>* one (v - int w))\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5598
        apply (simp add:reps.simps ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5599
        apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5600
        apply (rule tm.pre_condI, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5601
        apply (subgoal_tac " u + int (k - Suc w) = v - (1 + int w)", simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5602
        defer
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5603
        apply simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5604
        apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5605
        by (simp add:sep_conj_ac, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5606
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5607
    apply (hstep this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5608
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5609
      fix i' j'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5610
      have "\<lbrace>st i' \<and>* ps (v - (1 + int w)) \<and>* reps u (v - (1 + int w)) [k - Suc w]\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5611
               i' :[ if_zero j ]: j'
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5612
            \<lbrace>st j' \<and>* ps (v - (1 + int w)) \<and>* reps u (v - (1 + int w)) [k - Suc w]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5613
        apply (simp add:reps.simps ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5614
        apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5615
        apply (rule tm.pre_condI, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5616
        apply (subgoal_tac " u + int (k - Suc w) = v - (1 + int w)", simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5617
        defer
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5618
        apply simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5619
        by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5620
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5621
    apply (hstep this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5622
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5623
      fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5624
      have "\<lbrace>st i \<and>* ps (v - (1 + int w)) \<and>*  reps u (v - (1 + int w)) [k - Suc w]\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5625
                i :[ move_right ]: j 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5626
            \<lbrace>st j \<and>* ps (v - int w) \<and>*  reps u (v - (1 + int w)) [k - Suc w] \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5627
        apply (simp add:reps.simps ones_rev)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5628
        apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5629
        apply (rule tm.pre_condI, clarsimp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5630
        apply (subgoal_tac " u + int (k - Suc w) =  v - (1 + int w)", simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5631
        defer
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5632
        apply simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5633
        by hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5634
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5635
    apply (hsteps this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5636
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5637
      fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5638
      have "\<lbrace>st i \<and>* ps (v - int w) \<and>*  one (v + 2) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5639
                         zero (v - int w) \<and>* zeros (v - int w + 1) (v + 1)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5640
                 i :[right_until_one]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5641
            \<lbrace>st j \<and>* ps (v + 2) \<and>*  one (v + 2) \<and>*  zero (v - int w) \<and>* zeros (v - int w + 1) (v + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5642
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5643
          have "(zero (v - int w) \<and>* zeros (v - int w + 1) (v + 1)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5644
                    (zeros (v - int w) (v + 1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5645
            by (simp add:zeros_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5646
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5647
        apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5648
        by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5649
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5650
    apply (hstep this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5651
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5652
      from Suc(2, 3) have "w < k" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5653
      hence "(zeros (v + 3 + int w) (2 + (v + int k))) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5654
                  (zero (v + 3 + int w) \<and>* zeros (4 + (v + int w)) (2 + (v + int k)))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5655
        by (simp add:zeros_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5656
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5657
    apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5658
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5659
      fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5660
      have "\<lbrace>st i \<and>* ps (v + 2) \<and>* zero (v + 3 + int w) \<and>* zeros (4 + (v + int w)) (2 + (v + int k)) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5661
                                                        one (v + 2) \<and>* ones (v + 3) (v + 2 + int w)\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5662
                i :[right_until_zero]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5663
            \<lbrace>st j \<and>* ps (v + 3 + int w) \<and>* zero (v + 3 + int w) \<and>* zeros (4 + (v + int w)) (2 + (v + int k)) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5664
                                                        one (v + 2) \<and>* ones (v + 3) (v + 2 + int w)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5665
        my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5666
          have "(one (v + 2) \<and>* ones (v + 3) (v + 2 + int w)) =
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5667
                (ones (v + 2) (v + 2 + int w))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5668
            by (simp add:ones_simps, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5669
        my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5670
        apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5671
        by hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5672
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5673
    apply (hsteps this, simp only:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5674
    apply (sep_cancel+, simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5675
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5676
      fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5677
      assume "(zero (v - int w) \<and>* zeros (v - int w + 1) (v + 1)) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5678
      hence "zeros (v - int w) (v + 1) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5679
        by (simp add:zeros_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5680
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5681
    apply (fwd this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5682
    my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5683
      fix s
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5684
      assume "(one (v + 3 + int w) \<and>* ones (v + 3) (v + 2 + int w)) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5685
      hence "ones (v + 3) (3 + (v + int w)) s"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5686
        by (simp add:ones_rev sep_conj_ac, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5687
    my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5688
    apply (fwd this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5689
    by (simp add:sep_conj_ac, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5690
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5691
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5692
definition "cinit = (right_until_zero; move_right; write_one)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5693
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5694
definition "copy = (cinit; cmove; move_right; move_right; right_until_one; move_left; move_left; cfill_until_one)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5695
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5696
lemma hoare_copy:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5697
  shows
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5698
   "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u v [k] \<and>* zero (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5699
                                                     zeros (v + 2) (v + int(reps_len [k]) + 1)\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5700
                                  i :[copy]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5701
    \<lbrace>st j \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u v [k] \<and>* zero (v + 1) \<and>* 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5702
                                                      reps (v + 2) (v + int (reps_len [k]) + 1) [k]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5703
  apply (unfold copy_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5704
  my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5705
    fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5706
    have 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5707
       "\<lbrace>st i \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1) \<and>* zeros (v + 2) (v + int(reps_len [k]) + 1)\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5708
                      i :[cinit]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5709
        \<lbrace>st j \<and>* ps (v + 2) \<and>* reps u v [k] \<and>* zero (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5710
                                           one (v + 2) \<and>* zeros (v + 3) (v + int(reps_len [k]) + 1)\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5711
      apply (unfold cinit_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5712
      apply (simp add:reps.simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5713
      apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5714
      apply (rule tm.pre_condI, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5715
      apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5716
      apply (simp add:sep_conj_ac)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5717
      my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5718
        have "(zeros (u + int k + 2) (u + int k + int (reps_len [k]) + 1)) = 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5719
              (zero (u + int k + 2) \<and>*  zeros (u + int k + 3) (u + int k + int (reps_len [k]) + 1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5720
          by (smt reps_len_sg zeros_step_simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5721
      my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5722
      apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5723
      apply hstep
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5724
      by (simp add:sep_conj_ac, sep_cancel+, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5725
  my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5726
  apply (hstep this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5727
  apply (rule_tac p = "st j' \<and>* ps (v + 2) \<and>* reps u v [k] \<and>* zero (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5728
          one (v + 2) \<and>* zeros (v + 3) (v + int (reps_len [k]) + 1) \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5729
            <(v = u + int (reps_len [k]) - 1)>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5730
    " in tm.pre_stren)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5731
  my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5732
    fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5733
    from hoare_cmove[where w = 0 and k = k and i = i and j = j and v = v and u = u]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5734
    have "\<lbrace>st i \<and>* ps (v + 2) \<and>* zero (u - 1) \<and>* reps u v [k] \<and>* zero (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5735
                                            one (v + 2) \<and>* zeros (v + 3) (v + int(reps_len [k]) + 1)\<rbrace>
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5736
                      i :[cmove]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5737
          \<lbrace>st j \<and>* ps (u - 1) \<and>* zero (u - 1) \<and>* reps u u [0] \<and>* zeros (u + 1) (v + 1) \<and>*
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5738
                                                       reps (v + 2) (v + int (reps_len [k]) + 1) [k]\<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5739
      by (auto simp:ones_simps zeros_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5740
  my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5741
  apply (hstep this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5742
  apply (hstep, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5743
  my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5744
    have "reps u u [0] = one u" by (simp add:reps.simps ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5745
  my_block_end my_note eq_repsz = this
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5746
  apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5747
  apply (hstep)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5748
  apply (subst reps.simps, simp add: ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5749
  apply hsteps
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5750
  apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5751
  apply (rule tm.pre_condI, simp del:zeros.simps zeros_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5752
  apply (thin_tac "int (reps_len [k]) = 1 + int k \<and> v = u + int (reps_len [k]) - 1")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5753
  my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5754
    have "(zeros (u + 1) (u + int k + 1)) = (zeros (u + 1) (u + int k) \<and>* zero (u + int k + 1))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5755
      by (simp only:zeros_rev, smt)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5756
  my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5757
  apply (unfold this)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5758
  apply (hstep, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5759
  my_block
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5760
    fix i j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5761
    from hoare_cfill_until_one[where v = "u + int k" and u = "u + 1"]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5762
    have "\<lbrace>st i \<and>* ps (u + int k) \<and>* one u \<and>* zeros (u + 1) (u + int k)\<rbrace> 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5763
              i :[ cfill_until_one ]: j
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5764
          \<lbrace>st j \<and>* ps u \<and>* ones u (u + int k) \<rbrace>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5765
      by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5766
  my_block_end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5767
  apply (hstep this, simp add:sep_conj_ac reps.simps ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5768
  apply (simp add:sep_conj_ac reps.simps ones_simps)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5769
  apply (subst sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5770
  apply (subst (asm) sep_conj_cond)+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5771
  apply (rule condI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5772
  apply (erule condE, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5773
  apply (simp add: reps_len_def reps_sep_len_def reps_ctnt_len_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5774
  apply (sep_cancel+)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5775
  by (erule condE, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5776
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  5777
end