src/HOL/Imperative_HOL/Relational.thy
author haftmann
Tue, 06 Jul 2010 09:21:13 +0200
changeset 37724 6607ccf77946
parent 37719 271ecd4fb9f9
child 37725 6d28a2aea936
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     1
theory Relational 
28744
9257bb7bcd2d removed Assert.thy
haftmann
parents: 28145
diff changeset
     2
imports Array Ref
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     3
begin
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     4
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     5
section{* Definition of the Relational framework *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     6
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     7
text {* The crel predicate states that when a computation c runs with the heap h
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     8
  will result in return value r and a heap h' (if no exception occurs). *}  
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     9
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    10
definition crel :: "'a Heap \<Rightarrow> heap \<Rightarrow> heap \<Rightarrow> 'a \<Rightarrow> bool"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    11
where
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
    12
  crel_def': "crel c h h' r \<longleftrightarrow> Heap_Monad.execute c h = Some (r, h')"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    13
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    14
lemma crel_def: -- FIXME
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
    15
  "crel c h h' r \<longleftrightarrow> Some (r, h') = Heap_Monad.execute c h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    16
  unfolding crel_def' by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    17
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    18
lemma crel_deterministic: "\<lbrakk> crel f h h' a; crel f h h'' b \<rbrakk> \<Longrightarrow> (a = b) \<and> (h' = h'')"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    19
unfolding crel_def' by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    20
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    21
section {* Elimination rules *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    22
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    23
text {* For all commands, we define simple elimination rules. *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    24
(* FIXME: consumes 1 necessary ?? *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    25
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    26
subsection {* Elimination rules for basic monadic commands *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    27
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    28
lemma crelE[consumes 1]:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    29
  assumes "crel (f >>= g) h h'' r'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    30
  obtains h' r where "crel f h h' r" "crel (g r) h' h'' r'"
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
    31
  using assms by (auto simp add: crel_def bindM_def split: option.split_asm)
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    32
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    33
lemma crelE'[consumes 1]:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    34
  assumes "crel (f >> g) h h'' r'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    35
  obtains h' r where "crel f h h' r" "crel g h' h'' r'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    36
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    37
  by (elim crelE) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    38
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    39
lemma crel_return[consumes 1]:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    40
  assumes "crel (return x) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    41
  obtains "r = x" "h = h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    42
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    43
  unfolding crel_def return_def by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    44
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    45
lemma crel_raise[consumes 1]:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    46
  assumes "crel (raise x) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    47
  obtains "False"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    48
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    49
  unfolding crel_def raise_def by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    50
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    51
lemma crel_if:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    52
  assumes "crel (if c then t else e) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    53
  obtains "c" "crel t h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    54
        | "\<not>c" "crel e h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    55
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    56
  unfolding crel_def by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    57
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    58
lemma crel_option_case:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    59
  assumes "crel (case x of None \<Rightarrow> n | Some y \<Rightarrow> s y) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    60
  obtains "x = None" "crel n h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    61
        | y where "x = Some y" "crel (s y) h h' r" 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    62
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    63
  unfolding crel_def by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    64
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    65
lemma crel_mapM:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    66
  assumes "crel (mapM f xs) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    67
  assumes "\<And>h h'. P f [] h h' []"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    68
  assumes "\<And>h h1 h' x xs y ys. \<lbrakk> crel (f x) h h1 y; crel (mapM f xs) h1 h' ys; P f xs h1 h' ys \<rbrakk> \<Longrightarrow> P f (x#xs) h h' (y#ys)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    69
  shows "P f xs h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    70
using assms(1)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    71
proof (induct xs arbitrary: h h' r)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    72
  case Nil  with assms(2) show ?case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    73
    by (auto elim: crel_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    74
next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    75
  case (Cons x xs)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    76
  from Cons(2) obtain h1 y ys where crel_f: "crel (f x) h h1 y"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    77
    and crel_mapM: "crel (mapM f xs) h1 h' ys"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    78
    and r_def: "r = y#ys"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
    79
    unfolding mapM.simps
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    80
    by (auto elim!: crelE crel_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    81
  from Cons(1)[OF crel_mapM] crel_mapM crel_f assms(3) r_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    82
  show ?case by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    83
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    84
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    85
lemma crel_heap:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    86
  assumes "crel (Heap_Monad.heap f) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    87
  obtains "h' = snd (f h)" "r = fst (f h)"
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
    88
  using assms by (cases "f h") (simp add: crel_def)
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
    89
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    90
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    91
subsection {* Elimination rules for array commands *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    92
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    93
lemma crel_length:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
    94
  assumes "crel (len a) h h' r"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
    95
  obtains "h = h'" "r = Array.length a h'"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    96
  using assms
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
    97
  unfolding Array.len_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    98
  by (elim crel_heap) simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    99
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   100
(* Strong version of the lemma for this operation is missing *) 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   101
lemma crel_new_weak:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   102
  assumes "crel (Array.new n v) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   103
  obtains "get_array r h' = List.replicate n v"
37716
24bb91462892 remove primitive operation Heap.array in favour of Heap.array_of_list
haftmann
parents: 37709
diff changeset
   104
  using assms unfolding Array.new_def
24bb91462892 remove primitive operation Heap.array in favour of Heap.array_of_list
haftmann
parents: 37709
diff changeset
   105
  by (elim crel_heap) (auto simp: array_def Let_def split_def)
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   106
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   107
lemma crel_nth[consumes 1]:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   108
  assumes "crel (nth a i) h h' r"
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   109
  obtains "r = (get_array a h) ! i" "h = h'" "i < Array.length a h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   110
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   111
  unfolding nth_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   112
  by (auto elim!: crelE crel_if crel_raise crel_length crel_heap)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   113
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   114
lemma crel_upd[consumes 1]:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   115
  assumes "crel (upd i v a) h h' r"
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   116
  obtains "r = a" "h' = Array.change a i v h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   117
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   118
  unfolding upd_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   119
  by (elim crelE crel_if crel_return crel_raise
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   120
    crel_length crel_heap) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   121
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   122
(* Strong version of the lemma for this operation is missing *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   123
lemma crel_of_list_weak:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   124
  assumes "crel (Array.of_list xs) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   125
  obtains "get_array r h' = xs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   126
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   127
  unfolding of_list_def 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   128
  by (elim crel_heap) (simp add:get_array_init_array_list)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   129
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   130
lemma crel_map_entry:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   131
  assumes "crel (Array.map_entry i f a) h h' r"
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   132
  obtains "r = a" "h' = Array.change a i (f (get_array a h ! i)) h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   133
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   134
  unfolding Array.map_entry_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   135
  by (elim crelE crel_upd crel_nth) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   136
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   137
lemma crel_swap:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   138
  assumes "crel (Array.swap i x a) h h' r"
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   139
  obtains "r = get_array a h ! i" "h' = Array.change a i x h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   140
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   141
  unfolding Array.swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   142
  by (elim crelE crel_upd crel_nth crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   143
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   144
(* Strong version of the lemma for this operation is missing *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   145
lemma crel_make_weak:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   146
  assumes "crel (Array.make n f) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   147
  obtains "i < n \<Longrightarrow> get_array r h' ! i = f i"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   148
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   149
  unfolding Array.make_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   150
  by (elim crel_of_list_weak) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   151
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   152
lemma upt_conv_Cons':
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   153
  assumes "Suc a \<le> b"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   154
  shows "[b - Suc a..<b] = (b - Suc a)#[b - a..<b]"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   155
proof -
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   156
  from assms have l: "b - Suc a < b" by arith
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   157
  from assms have "Suc (b - Suc a) = b - a" by arith
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   158
  with l show ?thesis by (simp add: upt_conv_Cons)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   159
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   160
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   161
lemma crel_mapM_nth:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   162
  assumes
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   163
  "crel (mapM (Array.nth a) [Array.length a h - n..<Array.length a h]) h h' xs"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   164
  assumes "n \<le> Array.length a h"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   165
  shows "h = h' \<and> xs = drop (Array.length a h - n) (get_array a h)"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   166
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   167
proof (induct n arbitrary: xs h h')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   168
  case 0 thus ?case
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   169
    by (auto elim!: crel_return simp add: Array.length_def)
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   170
next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   171
  case (Suc n)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   172
  from Suc(3) have "[Array.length a h - Suc n..<Array.length a h] = (Array.length a h - Suc n)#[Array.length a h - n..<Array.length a h]"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   173
    by (simp add: upt_conv_Cons')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   174
  with Suc(2) obtain r where
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   175
    crel_mapM: "crel (mapM (Array.nth a) [Array.length a h - n..<Array.length a h]) h h' r"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   176
    and xs_def: "xs = get_array a h ! (Array.length a h - Suc n) # r"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   177
    by (auto elim!: crelE crel_nth crel_return)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   178
  from Suc(3) have "Array.length a h - n = Suc (Array.length a h - Suc n)" 
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   179
    by arith
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   180
  with Suc.hyps[OF crel_mapM] xs_def show ?case
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   181
    unfolding Array.length_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   182
    by (auto simp add: nth_drop')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   183
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   184
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   185
lemma crel_freeze:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   186
  assumes "crel (Array.freeze a) h h' xs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   187
  obtains "h = h'" "xs = get_array a h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   188
proof 
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   189
  from assms have "crel (mapM (Array.nth a) [0..<Array.length a h]) h h' xs"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   190
    unfolding freeze_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   191
    by (auto elim: crelE crel_length)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   192
  hence "crel (mapM (Array.nth a) [(Array.length a h - Array.length a h)..<Array.length a h]) h h' xs"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   193
    by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   194
  from crel_mapM_nth[OF this] show "h = h'" and "xs = get_array a h" by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   195
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   196
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   197
lemma crel_mapM_map_entry_remains:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   198
  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) h h' r"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   199
  assumes "i < Array.length a h - n"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   200
  shows "get_array a h ! i = get_array a h' ! i"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   201
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   202
proof (induct n arbitrary: h h' r)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   203
  case 0
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   204
  thus ?case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   205
    by (auto elim: crel_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   206
next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   207
  case (Suc n)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   208
  let ?h1 = "Array.change a (Array.length a h - Suc n) (f (get_array a h ! (Array.length a h - Suc n))) h"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   209
  from Suc(3) have "[Array.length a h - Suc n..<Array.length a h] = (Array.length a h - Suc n)#[Array.length a h - n..<Array.length a h]"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   210
    by (simp add: upt_conv_Cons')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   211
  from Suc(2) this obtain r where
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   212
    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) ?h1 h' r"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   213
    by (auto simp add: elim!: crelE crel_map_entry crel_return)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   214
  have length_remains: "Array.length a ?h1 = Array.length a h" by simp
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   215
  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a ?h1 - n..<Array.length a ?h1]) ?h1 h' r"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   216
    by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   217
  from Suc(1)[OF this] length_remains Suc(3) show ?case by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   218
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   219
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   220
lemma crel_mapM_map_entry_changes:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   221
  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) h h' r"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   222
  assumes "n \<le> Array.length a h"  
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   223
  assumes "i \<ge> Array.length a h - n"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   224
  assumes "i < Array.length a h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   225
  shows "get_array a h' ! i = f (get_array a h ! i)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   226
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   227
proof (induct n arbitrary: h h' r)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   228
  case 0
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   229
  thus ?case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   230
    by (auto elim!: crel_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   231
next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   232
  case (Suc n)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   233
  let ?h1 = "Array.change a (Array.length a h - Suc n) (f (get_array a h ! (Array.length a h - Suc n))) h"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   234
  from Suc(3) have "[Array.length a h - Suc n..<Array.length a h] = (Array.length a h - Suc n)#[Array.length a h - n..<Array.length a h]"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   235
    by (simp add: upt_conv_Cons')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   236
  from Suc(2) this obtain r where
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   237
    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) ?h1 h' r"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   238
    by (auto simp add: elim!: crelE crel_map_entry crel_return)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   239
  have length_remains: "Array.length a ?h1 = Array.length a h" by simp
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   240
  from Suc(3) have less: "Array.length a h - Suc n < Array.length a h - n" by arith
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   241
  from Suc(3) have less2: "Array.length a h - Suc n < Array.length a h" by arith
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   242
  from Suc(4) length_remains have cases: "i = Array.length a ?h1 - Suc n \<or> i \<ge> Array.length a ?h1 - n" by arith
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   243
  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a ?h1 - n..<Array.length a ?h1]) ?h1 h' r"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   244
    by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   245
  from Suc(1)[OF this] cases Suc(3) Suc(5) length_remains
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   246
    crel_mapM_map_entry_remains[OF this, of "Array.length a h - Suc n", symmetric] less less2
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   247
  show ?case
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   248
    by (auto simp add: nth_list_update_eq Array.length_def)
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   249
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   250
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   251
lemma crel_mapM_map_entry_length:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   252
  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) h h' r"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   253
  assumes "n \<le> Array.length a h"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   254
  shows "Array.length a h' = Array.length a h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   255
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   256
proof (induct n arbitrary: h h' r)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   257
  case 0
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   258
  thus ?case by (auto elim!: crel_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   259
next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   260
  case (Suc n)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   261
  let ?h1 = "Array.change a (Array.length a h - Suc n) (f (get_array a h ! (Array.length a h - Suc n))) h"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   262
  from Suc(3) have "[Array.length a h - Suc n..<Array.length a h] = (Array.length a h - Suc n)#[Array.length a h - n..<Array.length a h]"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   263
    by (simp add: upt_conv_Cons')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   264
  from Suc(2) this obtain r where 
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   265
    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) ?h1 h' r"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   266
    by (auto elim!: crelE crel_map_entry crel_return)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   267
  have length_remains: "Array.length a ?h1 = Array.length a h" by simp
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   268
  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a ?h1 - n..<Array.length a ?h1]) ?h1 h' r"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   269
    by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   270
  from Suc(1)[OF this] length_remains Suc(3) show ?case by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   271
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   272
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   273
lemma crel_mapM_map_entry:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   274
assumes "crel (mapM (\<lambda>n. map_entry n f a) [0..<Array.length a h]) h h' r"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   275
  shows "get_array a h' = List.map f (get_array a h)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   276
proof -
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   277
  from assms have "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - Array.length a h..<Array.length a h]) h h' r" by simp
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   278
  from crel_mapM_map_entry_length[OF this]
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   279
  crel_mapM_map_entry_changes[OF this] show ?thesis
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   280
    unfolding Array.length_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   281
    by (auto intro: nth_equalityI)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   282
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   283
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   284
lemma crel_map_weak:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   285
  assumes crel_map: "crel (Array.map f a) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   286
  obtains "r = a" "get_array a h' = List.map f (get_array a h)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   287
proof
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   288
  from assms crel_mapM_map_entry show  "get_array a h' = List.map f (get_array a h)"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   289
    unfolding Array.map_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   290
    by (fastsimp elim!: crelE crel_length crel_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   291
  from assms show "r = a"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   292
    unfolding Array.map_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   293
    by (elim crelE crel_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   294
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   295
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   296
subsection {* Elimination rules for reference commands *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   297
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   298
(* TODO:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   299
maybe introduce a new predicate "extends h' h x"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   300
which means h' extends h with a new reference x.
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   301
Then crel_new: would be
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   302
assumes "crel (Ref.new v) h h' x"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   303
obtains "get_ref x h' = v"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   304
and "extends h' h x"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   305
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   306
and we would need further rules for extends:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   307
extends h' h x \<Longrightarrow> \<not> ref_present x h
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   308
extends h' h x \<Longrightarrow>  ref_present x h'
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   309
extends h' h x \<Longrightarrow> ref_present y h \<Longrightarrow> ref_present y h'
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   310
extends h' h x \<Longrightarrow> ref_present y h \<Longrightarrow> get_ref y h = get_ref y h'
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   311
extends h' h x \<Longrightarrow> lim h' = Suc (lim h)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   312
*)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   313
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   314
lemma crel_Ref_new:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   315
  assumes "crel (Ref.new v) h h' x"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   316
  obtains "get_ref x h' = v"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   317
  and "\<not> ref_present x h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   318
  and "ref_present x h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   319
  and "\<forall>y. ref_present y h \<longrightarrow> get_ref y h = get_ref y h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   320
 (* and "lim h' = Suc (lim h)" *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   321
  and "\<forall>y. ref_present y h \<longrightarrow> ref_present y h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   322
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   323
  unfolding Ref.new_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   324
  apply (elim crel_heap)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   325
  unfolding Ref.ref_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   326
  apply (simp add: Let_def)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   327
  unfolding ref_present_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   328
  apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   329
  unfolding get_ref_def set_ref_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   330
  apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   331
  done
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   332
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   333
lemma crel_lookup:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   334
  assumes "crel (!r') h h' r"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   335
  obtains "h = h'" "r = get_ref r' h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   336
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   337
unfolding Ref.lookup_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   338
by (auto elim: crel_heap)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   339
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   340
lemma crel_update:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   341
  assumes "crel (r' := v) h h' r"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   342
  obtains "h' = set_ref r' v h" "r = ()"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   343
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   344
unfolding Ref.update_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   345
by (auto elim: crel_heap)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   346
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   347
lemma crel_change:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   348
  assumes "crel (Ref.change f r') h h' r"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   349
  obtains "h' = set_ref r' (f (get_ref r' h)) h" "r = f (get_ref r' h)"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   350
using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   351
unfolding Ref.change_def Let_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   352
by (auto elim!: crelE crel_lookup crel_update crel_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   353
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   354
subsection {* Elimination rules for the assert command *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   355
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   356
lemma crel_assert[consumes 1]:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   357
  assumes "crel (assert P x) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   358
  obtains "P x" "r = x" "h = h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   359
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   360
  unfolding assert_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   361
  by (elim crel_if crel_return crel_raise) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   362
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   363
lemma crel_assert_eq: "(\<And>h h' r. crel f h h' r \<Longrightarrow> P r) \<Longrightarrow> f \<guillemotright>= assert P = f"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   364
unfolding crel_def bindM_def Let_def assert_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   365
  raise_def return_def prod_case_beta
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   366
apply (cases f)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   367
apply simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   368
apply (simp add: expand_fun_eq split_def)
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   369
apply (auto split: option.split)
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   370
apply (erule_tac x="x" in meta_allE)
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   371
apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   372
done
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   373
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   374
section {* Introduction rules *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   375
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   376
subsection {* Introduction rules for basic monadic commands *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   377
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   378
lemma crelI:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   379
  assumes "crel f h h' r" "crel (g r) h' h'' r'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   380
  shows "crel (f >>= g) h h'' r'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   381
  using assms by (simp add: crel_def' bindM_def)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   382
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   383
lemma crelI':
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   384
  assumes "crel f h h' r" "crel g h' h'' r'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   385
  shows "crel (f >> g) h h'' r'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   386
  using assms by (intro crelI) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   387
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   388
lemma crel_returnI:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   389
  shows "crel (return x) h h x"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   390
  unfolding crel_def return_def by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   391
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   392
lemma crel_raiseI:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   393
  shows "\<not> (crel (raise x) h h' r)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   394
  unfolding crel_def raise_def by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   395
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   396
lemma crel_ifI:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   397
  assumes "c \<longrightarrow> crel t h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   398
  "\<not>c \<longrightarrow> crel e h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   399
  shows "crel (if c then t else e) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   400
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   401
  unfolding crel_def by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   402
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   403
lemma crel_option_caseI:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   404
  assumes "\<And>y. x = Some y \<Longrightarrow> crel (s y) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   405
  assumes "x = None \<Longrightarrow> crel n h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   406
  shows "crel (case x of None \<Rightarrow> n | Some y \<Rightarrow> s y) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   407
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   408
by (auto split: option.split)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   409
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   410
lemma crel_heapI:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   411
  shows "crel (Heap_Monad.heap f) h (snd (f h)) (fst (f h))"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   412
  by (simp add: crel_def apfst_def split_def prod_fun_def)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   413
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   414
lemma crel_heapI':
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   415
  assumes "h' = snd (f h)" "r = fst (f h)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   416
  shows "crel (Heap_Monad.heap f) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   417
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   418
  by (simp add: crel_def split_def apfst_def prod_fun_def)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   419
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   420
lemma crelI2:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   421
  assumes "\<exists>h' rs'. crel f h h' rs' \<and> (\<exists>h'' rs. crel (g rs') h' h'' rs)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   422
  shows "\<exists>h'' rs. crel (f\<guillemotright>= g) h h'' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   423
  oops
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   424
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   425
lemma crel_ifI2:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   426
  assumes "c \<Longrightarrow> \<exists>h' r. crel t h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   427
  "\<not> c \<Longrightarrow> \<exists>h' r. crel e h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   428
  shows "\<exists> h' r. crel (if c then t else e) h h' r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   429
  oops
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   430
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   431
subsection {* Introduction rules for array commands *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   432
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   433
lemma crel_lengthI:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   434
  shows "crel (Array.len a) h h (Array.length a h)"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   435
  unfolding len_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   436
  by (rule crel_heapI') auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   437
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   438
(* thm crel_newI for Array.new is missing *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   439
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   440
lemma crel_nthI:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   441
  assumes "i < Array.length a h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   442
  shows "crel (nth a i) h h ((get_array a h) ! i)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   443
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   444
  unfolding nth_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   445
  by (auto intro!: crelI crel_ifI crel_raiseI crel_lengthI crel_heapI')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   446
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   447
lemma crel_updI:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   448
  assumes "i < Array.length a h"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   449
  shows "crel (upd i v a) h (Array.change a i v h) a"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   450
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   451
  unfolding upd_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   452
  by (auto intro!: crelI crel_ifI crel_returnI crel_raiseI
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   453
    crel_lengthI crel_heapI')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   454
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   455
(* thm crel_of_listI is missing *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   456
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   457
(* thm crel_map_entryI is missing *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   458
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   459
(* thm crel_swapI is missing *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   460
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   461
(* thm crel_makeI is missing *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   462
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   463
(* thm crel_freezeI is missing *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   464
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   465
(* thm crel_mapI is missing *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   466
36057
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   467
subsubsection {* Introduction rules for reference commands *}
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   468
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   469
lemma crel_lookupI:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   470
  shows "crel (!r) h h (get_ref r h)"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   471
  unfolding lookup_def by (auto intro!: crel_heapI')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   472
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   473
lemma crel_updateI:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   474
  shows "crel (r := v) h (set_ref r v h) ()"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   475
  unfolding update_def by (auto intro!: crel_heapI')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   476
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   477
lemma crel_changeI: 
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   478
  shows "crel (Ref.change f r) h (set_ref r (f (get_ref r h)) h) (f (get_ref r h))"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   479
unfolding change_def Let_def by (auto intro!: crelI crel_returnI crel_lookupI crel_updateI)
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   480
36057
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   481
subsubsection {* Introduction rules for the assert command *}
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   482
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   483
lemma crel_assertI:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   484
  assumes "P x"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   485
  shows "crel (assert P x) h h x"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   486
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   487
  unfolding assert_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   488
  by (auto intro!: crel_ifI crel_returnI crel_raiseI)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   489
 
36057
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   490
subsection {* Induction rule for the MREC combinator *}
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   491
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   492
lemma MREC_induct:
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   493
  assumes "crel (MREC f g x) h h' r"
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   494
  assumes "\<And> x h h' r. crel (f x) h h' (Inl r) \<Longrightarrow> P x h h' r"
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   495
  assumes "\<And> x h h1 h2 h' s z r. crel (f x) h h1 (Inr s) \<Longrightarrow> crel (MREC f g s) h1 h2 z \<Longrightarrow> P s h1 h2 z
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   496
    \<Longrightarrow> crel (g x s z) h2 h' r \<Longrightarrow> P x h h' r"
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   497
  shows "P x h h' r"
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   498
proof (rule MREC_pinduct[OF assms(1)[unfolded crel_def, symmetric]])
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   499
  fix x h h1 h2 h' s z r
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   500
  assume "Heap_Monad.execute (f x) h = Some (Inr s, h1)"
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   501
    "Heap_Monad.execute (MREC f g s) h1 = Some (z, h2)"
36057
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   502
    "P s h1 h2 z"
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   503
    "Heap_Monad.execute (g x s z) h2 = Some (r, h')"
36057
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   504
  from assms(3)[unfolded crel_def, OF this(1)[symmetric] this(2)[symmetric] this(3) this(4)[symmetric]]
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   505
  show "P x h h' r" .
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   506
next
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   507
qed (auto simp add: assms(2)[unfolded crel_def])
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   508
ca6610908ae9 adding MREC induction rule in Imperative HOL
bulwahn
parents: 29399
diff changeset
   509
section {* Definition of the noError predicate *}
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   510
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   511
text {* We add a simple definitional setting for crel intro rules
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   512
  where we only would like to show that the computation does not result in a exception for heap h,
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   513
  but we do not care about statements about the resulting heap and return value.*}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   514
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   515
definition noError :: "'a Heap \<Rightarrow> heap \<Rightarrow> bool"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   516
where
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   517
  "noError c h \<longleftrightarrow> (\<exists>r h'. Some (r, h') = Heap_Monad.execute c h)"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   518
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   519
lemma noError_def': -- FIXME
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   520
  "noError c h \<longleftrightarrow> (\<exists>r h'. Heap_Monad.execute c h = Some (r, h'))"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   521
  unfolding noError_def apply auto proof -
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   522
  fix r h'
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   523
  assume "Some (r, h') = Heap_Monad.execute c h"
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   524
  then have "Heap_Monad.execute c h = Some (r, h')" ..
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   525
  then show "\<exists>r h'. Heap_Monad.execute c h = Some (r, h')" by blast
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   526
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   527
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   528
subsection {* Introduction rules for basic monadic commands *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   529
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   530
lemma noErrorI:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   531
  assumes "noError f h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   532
  assumes "\<And>h' r. crel f h h' r \<Longrightarrow> noError (g r) h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   533
  shows "noError (f \<guillemotright>= g) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   534
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   535
  by (auto simp add: noError_def' crel_def' bindM_def)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   536
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   537
lemma noErrorI':
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   538
  assumes "noError f h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   539
  assumes "\<And>h' r. crel f h h' r \<Longrightarrow> noError g h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   540
  shows "noError (f \<guillemotright> g) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   541
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   542
  by (auto simp add: noError_def' crel_def' bindM_def)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   543
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   544
lemma noErrorI2:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   545
"\<lbrakk>crel f h h' r ; noError f h; noError (g r) h'\<rbrakk>
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   546
\<Longrightarrow> noError (f \<guillemotright>= g) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   547
by (auto simp add: noError_def' crel_def' bindM_def)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   548
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   549
lemma noError_return: 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   550
  shows "noError (return x) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   551
  unfolding noError_def return_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   552
  by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   553
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   554
lemma noError_if:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   555
  assumes "c \<Longrightarrow> noError t h" "\<not> c \<Longrightarrow> noError e h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   556
  shows "noError (if c then t else e) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   557
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   558
  unfolding noError_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   559
  by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   560
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   561
lemma noError_option_case:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   562
  assumes "\<And>y. x = Some y \<Longrightarrow> noError (s y) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   563
  assumes "noError n h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   564
  shows "noError (case x of None \<Rightarrow> n | Some y \<Rightarrow> s y) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   565
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   566
by (auto split: option.split)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   567
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   568
lemma noError_mapM: 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   569
assumes "\<forall>x \<in> set xs. noError (f x) h \<and> crel (f x) h h (r x)" 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   570
shows "noError (mapM f xs) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   571
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   572
proof (induct xs)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   573
  case Nil
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   574
  thus ?case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   575
    unfolding mapM.simps by (intro noError_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   576
next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   577
  case (Cons x xs)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   578
  thus ?case
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   579
    unfolding mapM.simps
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   580
    by (auto intro: noErrorI2[of "f x"] noErrorI noError_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   581
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   582
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   583
lemma noError_heap:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   584
  shows "noError (Heap_Monad.heap f) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   585
  by (simp add: noError_def' apfst_def prod_fun_def split_def)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   586
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   587
subsection {* Introduction rules for array commands *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   588
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   589
lemma noError_length:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   590
  shows "noError (Array.len a) h"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   591
  unfolding len_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   592
  by (intro noError_heap)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   593
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   594
lemma noError_new:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   595
  shows "noError (Array.new n v) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   596
unfolding Array.new_def by (intro noError_heap)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   597
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   598
lemma noError_upd:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   599
  assumes "i < Array.length a h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   600
  shows "noError (Array.upd i v a) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   601
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   602
  unfolding upd_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   603
  by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   604
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   605
lemma noError_nth:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   606
assumes "i < Array.length a h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   607
  shows "noError (Array.nth a i) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   608
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   609
  unfolding nth_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   610
  by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   611
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   612
lemma noError_of_list:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   613
  shows "noError (of_list ls) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   614
  unfolding of_list_def by (rule noError_heap)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   615
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   616
lemma noError_map_entry:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   617
  assumes "i < Array.length a h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   618
  shows "noError (map_entry i f a) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   619
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   620
  unfolding map_entry_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   621
  by (auto elim: crel_nth intro!: noErrorI noError_nth noError_upd)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   622
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   623
lemma noError_swap:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   624
  assumes "i < Array.length a h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   625
  shows "noError (swap i x a) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   626
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   627
  unfolding swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   628
  by (auto elim: crel_nth intro!: noErrorI noError_return noError_nth noError_upd)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   629
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   630
lemma noError_make:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   631
  shows "noError (make n f) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   632
  unfolding make_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   633
  by (auto intro: noError_of_list)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   634
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   635
(*TODO: move to HeapMonad *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   636
lemma mapM_append:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   637
  "mapM f (xs @ ys) = mapM f xs \<guillemotright>= (\<lambda>xs. mapM f ys \<guillemotright>= (\<lambda>ys. return (xs @ ys)))"
37709
70fafefbcc98 simplified representation of monad type
haftmann
parents: 36057
diff changeset
   638
  by (induct xs) simp_all
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   639
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   640
lemma noError_freeze:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   641
  shows "noError (freeze a) h"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   642
unfolding freeze_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   643
by (auto intro!: noErrorI noError_length noError_mapM[of _ _ _ "\<lambda>x. get_array a h ! x"]
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   644
  noError_nth crel_nthI elim: crel_length)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   645
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   646
lemma noError_mapM_map_entry:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   647
  assumes "n \<le> Array.length a h"
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   648
  shows "noError (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   649
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   650
proof (induct n arbitrary: h)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   651
  case 0
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   652
  thus ?case by (auto intro: noError_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   653
next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   654
  case (Suc n)
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   655
  from Suc.prems have "[Array.length a h - Suc n..<Array.length a h] = (Array.length a h - Suc n)#[Array.length a h - n..<Array.length a h]"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   656
    by (simp add: upt_conv_Cons')
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   657
  with Suc.hyps[of "(Array.change a (Array.length a h - Suc n) (f (get_array a h ! (Array.length a h - Suc n))) h)"] Suc.prems show ?case
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   658
    by (auto simp add: intro!: noErrorI noError_return noError_map_entry elim: crel_map_entry)
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   659
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   660
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   661
lemma noError_map:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   662
  shows "noError (Array.map f a) h"
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   663
using noError_mapM_map_entry[of "Array.length a h" a h]
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   664
unfolding Array.map_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   665
by (auto intro: noErrorI noError_length noError_return elim!: crel_length)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   666
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   667
subsection {* Introduction rules for the reference commands *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   668
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   669
lemma noError_Ref_new:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   670
  shows "noError (Ref.new v) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   671
unfolding Ref.new_def by (intro noError_heap)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   672
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   673
lemma noError_lookup:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   674
  shows "noError (!r) h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   675
  unfolding lookup_def by (intro noError_heap)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   676
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   677
lemma noError_update:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   678
  shows "noError (r := v) h"
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   679
  unfolding update_def by (intro noError_heap)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   680
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   681
lemma noError_change:
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
   682
  shows "noError (Ref.change f r) h"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27656
diff changeset
   683
  unfolding Ref.change_def Let_def by (intro noErrorI noError_lookup noError_update noError_return)
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   684
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   685
subsection {* Introduction rules for the assert command *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   686
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   687
lemma noError_assert:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   688
  assumes "P x"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   689
  shows "noError (assert P x) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   690
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   691
  unfolding assert_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   692
  by (auto intro: noError_if noError_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   693
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   694
section {* Cumulative lemmas *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   695
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   696
lemmas crel_elim_all =
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   697
  crelE crelE' crel_return crel_raise crel_if crel_option_case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   698
  crel_length crel_new_weak crel_nth crel_upd crel_of_list_weak crel_map_entry crel_swap crel_make_weak crel_freeze crel_map_weak
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   699
  crel_Ref_new crel_lookup crel_update crel_change
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   700
  crel_assert
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   701
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   702
lemmas crel_intro_all =
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   703
  crelI crelI' crel_returnI crel_raiseI crel_ifI crel_option_caseI
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   704
  crel_lengthI (* crel_newI *) crel_nthI crel_updI (* crel_of_listI crel_map_entryI crel_swapI crel_makeI crel_freezeI crel_mapI *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   705
  (* crel_Ref_newI *) crel_lookupI crel_updateI crel_changeI
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   706
  crel_assert
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   707
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   708
lemmas noError_intro_all = 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   709
  noErrorI noErrorI' noError_return noError_if noError_option_case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   710
  noError_length noError_new noError_nth noError_upd noError_of_list noError_map_entry noError_swap noError_make noError_freeze noError_map
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   711
  noError_Ref_new noError_lookup noError_update noError_change
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   712
  noError_assert
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   713
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   714
end