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