| author | wenzelm | 
| Fri, 27 Jan 2012 21:56:29 +0100 | |
| changeset 46273 | 48cf461535cf | 
| parent 44890 | 22f665a2e91c | 
| child 47108 | 2a1953f0d20d | 
| permissions | -rw-r--r-- | 
| 36098 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35041diff
changeset | 1 | (* Title: HOL/Imperative_HOL/ex/Imperative_Quicksort.thy | 
| 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35041diff
changeset | 2 | Author: Lukas Bulwahn, TU Muenchen | 
| 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35041diff
changeset | 3 | *) | 
| 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35041diff
changeset | 4 | |
| 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35041diff
changeset | 5 | header {* An imperative implementation of Quicksort on arrays *}
 | 
| 30689 
b14b2cc4e25e
moved Imperative_HOL examples to Imperative_HOL/ex
 haftmann parents: 
29793diff
changeset | 6 | |
| 
b14b2cc4e25e
moved Imperative_HOL examples to Imperative_HOL/ex
 haftmann parents: 
29793diff
changeset | 7 | theory Imperative_Quicksort | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40671diff
changeset | 8 | imports | 
| 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40671diff
changeset | 9 | Imperative_HOL | 
| 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40671diff
changeset | 10 | Subarray | 
| 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40671diff
changeset | 11 | "~~/src/HOL/Library/Multiset" | 
| 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40671diff
changeset | 12 | "~~/src/HOL/Library/Efficient_Nat" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 13 | begin | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 14 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 15 | text {* We prove QuickSort correct in the Relational Calculus. *}
 | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 16 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 17 | definition swap :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 18 | where | 
| 37792 | 19 | "swap arr i j = | 
| 20 |      do {
 | |
| 37798 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 21 | x \<leftarrow> Array.nth arr i; | 
| 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 22 | y \<leftarrow> Array.nth arr j; | 
| 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 23 | Array.upd i y arr; | 
| 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 24 | Array.upd j x arr; | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 25 | return () | 
| 37792 | 26 | }" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 27 | |
| 40671 | 28 | lemma effect_swapI [effect_intros]: | 
| 37802 | 29 | assumes "i < Array.length h a" "j < Array.length h a" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 30 | "x = Array.get h a ! i" "y = Array.get h a ! j" | 
| 37798 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 31 | "h' = Array.update a j x (Array.update a i y h)" | 
| 40671 | 32 | shows "effect (swap a i j) h h' r" | 
| 33 | unfolding swap_def using assms by (auto intro!: effect_intros) | |
| 37771 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37758diff
changeset | 34 | |
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 35 | lemma swap_permutes: | 
| 40671 | 36 | assumes "effect (swap a i j) h h' rs" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 37 | shows "multiset_of (Array.get h' a) | 
| 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 38 | = multiset_of (Array.get h a)" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 39 | using assms | 
| 28145 | 40 | unfolding swap_def | 
| 40671 | 41 | by (auto simp add: Array.length_def multiset_of_swap dest: sym [of _ "h'"] elim!: effect_bindE effect_nthE effect_returnE effect_updE) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 42 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 43 | function part1 :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 44 | where | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 45 | "part1 a left right p = ( | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 46 | if (right \<le> left) then return right | 
| 37792 | 47 |      else do {
 | 
| 37798 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 48 | v \<leftarrow> Array.nth a left; | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 49 | (if (v \<le> p) then (part1 a (left + 1) right p) | 
| 37792 | 50 |                     else (do { swap a left right;
 | 
| 51 | part1 a left (right - 1) p })) | |
| 52 | })" | |
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 53 | by pat_completeness auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 54 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 55 | termination | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 56 | by (relation "measure (\<lambda>(_,l,r,_). r - l )") auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 57 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 58 | declare part1.simps[simp del] | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 59 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 60 | lemma part_permutes: | 
| 40671 | 61 | assumes "effect (part1 a l r p) h h' rs" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 62 | shows "multiset_of (Array.get h' a) | 
| 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 63 | = multiset_of (Array.get h a)" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 64 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 65 | proof (induct a l r p arbitrary: h h' rs rule:part1.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 66 | case (1 a l r p h h' rs) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 67 | thus ?case | 
| 28145 | 68 | unfolding part1.simps [of a l r p] | 
| 40671 | 69 | by (elim effect_bindE effect_ifE effect_returnE effect_nthE) (auto simp add: swap_permutes) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 70 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 71 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 72 | lemma part_returns_index_in_bounds: | 
| 40671 | 73 | assumes "effect (part1 a l r p) h h' rs" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 74 | assumes "l \<le> r" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 75 | shows "l \<le> rs \<and> rs \<le> r" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 76 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 77 | proof (induct a l r p arbitrary: h h' rs rule:part1.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 78 | case (1 a l r p h h' rs) | 
| 40671 | 79 | note cr = `effect (part1 a l r p) h h' rs` | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 80 | show ?case | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 81 | proof (cases "r \<le> l") | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 82 | case True (* Terminating case *) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 83 | with cr `l \<le> r` show ?thesis | 
| 28145 | 84 | unfolding part1.simps[of a l r p] | 
| 40671 | 85 | by (elim effect_bindE effect_ifE effect_returnE effect_nthE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 86 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 87 | case False (* recursive case *) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 88 | note rec_condition = this | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 89 | let ?v = "Array.get h a ! l" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 90 | show ?thesis | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 91 | proof (cases "?v \<le> p") | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 92 | case True | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 93 | with cr False | 
| 40671 | 94 | have rec1: "effect (part1 a (l + 1) r p) h h' rs" | 
| 28145 | 95 | unfolding part1.simps[of a l r p] | 
| 40671 | 96 | by (elim effect_bindE effect_nthE effect_ifE effect_returnE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 97 | from rec_condition have "l + 1 \<le> r" by arith | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 98 | from 1(1)[OF rec_condition True rec1 `l + 1 \<le> r`] | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 99 | show ?thesis by simp | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 100 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 101 | case False | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 102 | with rec_condition cr | 
| 40671 | 103 | obtain h1 where swp: "effect (swap a l r) h h1 ()" | 
| 104 | and rec2: "effect (part1 a l (r - 1) p) h1 h' rs" | |
| 28145 | 105 | unfolding part1.simps[of a l r p] | 
| 40671 | 106 | by (elim effect_bindE effect_nthE effect_ifE effect_returnE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 107 | from rec_condition have "l \<le> r - 1" by arith | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 108 | from 1(2) [OF rec_condition False rec2 `l \<le> r - 1`] show ?thesis by fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 109 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 110 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 111 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 112 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 113 | lemma part_length_remains: | 
| 40671 | 114 | assumes "effect (part1 a l r p) h h' rs" | 
| 37802 | 115 | shows "Array.length h a = Array.length h' a" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 116 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 117 | proof (induct a l r p arbitrary: h h' rs rule:part1.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 118 | case (1 a l r p h h' rs) | 
| 40671 | 119 | note cr = `effect (part1 a l r p) h h' rs` | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 120 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 121 | show ?case | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 122 | proof (cases "r \<le> l") | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 123 | case True (* Terminating case *) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 124 | with cr show ?thesis | 
| 28145 | 125 | unfolding part1.simps[of a l r p] | 
| 40671 | 126 | by (elim effect_bindE effect_ifE effect_returnE effect_nthE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 127 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 128 | case False (* recursive case *) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 129 | with cr 1 show ?thesis | 
| 28145 | 130 | unfolding part1.simps [of a l r p] swap_def | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 131 | by (auto elim!: effect_bindE effect_ifE effect_nthE effect_returnE effect_updE) fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 132 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 133 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 134 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 135 | lemma part_outer_remains: | 
| 40671 | 136 | assumes "effect (part1 a l r p) h h' rs" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 137 | shows "\<forall>i. i < l \<or> r < i \<longrightarrow> Array.get h (a::nat array) ! i = Array.get h' a ! i" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 138 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 139 | proof (induct a l r p arbitrary: h h' rs rule:part1.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 140 | case (1 a l r p h h' rs) | 
| 40671 | 141 | note cr = `effect (part1 a l r p) h h' rs` | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 142 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 143 | show ?case | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 144 | proof (cases "r \<le> l") | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 145 | case True (* Terminating case *) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 146 | with cr show ?thesis | 
| 28145 | 147 | unfolding part1.simps[of a l r p] | 
| 40671 | 148 | by (elim effect_bindE effect_ifE effect_returnE effect_nthE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 149 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 150 | case False (* recursive case *) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 151 | note rec_condition = this | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 152 | let ?v = "Array.get h a ! l" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 153 | show ?thesis | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 154 | proof (cases "?v \<le> p") | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 155 | case True | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 156 | with cr False | 
| 40671 | 157 | have rec1: "effect (part1 a (l + 1) r p) h h' rs" | 
| 28145 | 158 | unfolding part1.simps[of a l r p] | 
| 40671 | 159 | by (elim effect_bindE effect_nthE effect_ifE effect_returnE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 160 | from 1(1)[OF rec_condition True rec1] | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 161 | show ?thesis by fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 162 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 163 | case False | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 164 | with rec_condition cr | 
| 40671 | 165 | obtain h1 where swp: "effect (swap a l r) h h1 ()" | 
| 166 | and rec2: "effect (part1 a l (r - 1) p) h1 h' rs" | |
| 28145 | 167 | unfolding part1.simps[of a l r p] | 
| 40671 | 168 | by (elim effect_bindE effect_nthE effect_ifE effect_returnE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 169 | from swp rec_condition have | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 170 | "\<forall>i. i < l \<or> r < i \<longrightarrow> Array.get h a ! i = Array.get h1 a ! i" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 171 | unfolding swap_def | 
| 40671 | 172 | by (elim effect_bindE effect_nthE effect_updE effect_returnE) auto | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 173 | with 1(2) [OF rec_condition False rec2] show ?thesis by fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 174 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 175 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 176 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 177 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 178 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 179 | lemma part_partitions: | 
| 40671 | 180 | assumes "effect (part1 a l r p) h h' rs" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 181 | shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> Array.get h' (a::nat array) ! i \<le> p) | 
| 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 182 | \<and> (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> Array.get h' a ! i \<ge> p)" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 183 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 184 | proof (induct a l r p arbitrary: h h' rs rule:part1.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 185 | case (1 a l r p h h' rs) | 
| 40671 | 186 | note cr = `effect (part1 a l r p) h h' rs` | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 187 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 188 | show ?case | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 189 | proof (cases "r \<le> l") | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 190 | case True (* Terminating case *) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 191 | with cr have "rs = r" | 
| 28145 | 192 | unfolding part1.simps[of a l r p] | 
| 40671 | 193 | by (elim effect_bindE effect_ifE effect_returnE effect_nthE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 194 | with True | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 195 | show ?thesis by auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 196 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 197 | case False (* recursive case *) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 198 | note lr = this | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 199 | let ?v = "Array.get h a ! l" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 200 | show ?thesis | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 201 | proof (cases "?v \<le> p") | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 202 | case True | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 203 | with lr cr | 
| 40671 | 204 | have rec1: "effect (part1 a (l + 1) r p) h h' rs" | 
| 28145 | 205 | unfolding part1.simps[of a l r p] | 
| 40671 | 206 | by (elim effect_bindE effect_nthE effect_ifE effect_returnE) auto | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 207 | from True part_outer_remains[OF rec1] have a_l: "Array.get h' a ! l \<le> p" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 208 | by fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 209 | have "\<forall>i. (l \<le> i = (l = i \<or> Suc l \<le> i))" by arith | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 210 | with 1(1)[OF False True rec1] a_l show ?thesis | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 211 | by auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 212 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 213 | case False | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 214 | with lr cr | 
| 40671 | 215 | obtain h1 where swp: "effect (swap a l r) h h1 ()" | 
| 216 | and rec2: "effect (part1 a l (r - 1) p) h1 h' rs" | |
| 28145 | 217 | unfolding part1.simps[of a l r p] | 
| 40671 | 218 | by (elim effect_bindE effect_nthE effect_ifE effect_returnE) auto | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 219 | from swp False have "Array.get h1 a ! r \<ge> p" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 220 | unfolding swap_def | 
| 40671 | 221 | by (auto simp add: Array.length_def elim!: effect_bindE effect_nthE effect_updE effect_returnE) | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 222 | with part_outer_remains [OF rec2] lr have a_r: "Array.get h' a ! r \<ge> p" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 223 | by fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 224 | have "\<forall>i. (i \<le> r = (i = r \<or> i \<le> r - 1))" by arith | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 225 | with 1(2)[OF lr False rec2] a_r show ?thesis | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 226 | by auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 227 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 228 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 229 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 230 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 231 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 232 | fun partition :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 233 | where | 
| 37792 | 234 |   "partition a left right = do {
 | 
| 37798 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 235 | pivot \<leftarrow> Array.nth a right; | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 236 | middle \<leftarrow> part1 a left (right - 1) pivot; | 
| 37798 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 237 | v \<leftarrow> Array.nth a middle; | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 238 | m \<leftarrow> return (if (v \<le> pivot) then (middle + 1) else middle); | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 239 | swap a m right; | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 240 | return m | 
| 37792 | 241 | }" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 242 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 243 | declare partition.simps[simp del] | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 244 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 245 | lemma partition_permutes: | 
| 40671 | 246 | assumes "effect (partition a l r) h h' rs" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 247 | shows "multiset_of (Array.get h' a) | 
| 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 248 | = multiset_of (Array.get h a)" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 249 | proof - | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 250 | from assms part_permutes swap_permutes show ?thesis | 
| 28145 | 251 | unfolding partition.simps | 
| 40671 | 252 | by (elim effect_bindE effect_returnE effect_nthE effect_ifE effect_updE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 253 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 254 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 255 | lemma partition_length_remains: | 
| 40671 | 256 | assumes "effect (partition a l r) h h' rs" | 
| 37802 | 257 | shows "Array.length h a = Array.length h' a" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 258 | proof - | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 259 | from assms part_length_remains show ?thesis | 
| 28145 | 260 | unfolding partition.simps swap_def | 
| 40671 | 261 | by (elim effect_bindE effect_returnE effect_nthE effect_ifE effect_updE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 262 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 263 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 264 | lemma partition_outer_remains: | 
| 40671 | 265 | assumes "effect (partition a l r) h h' rs" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 266 | assumes "l < r" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 267 | shows "\<forall>i. i < l \<or> r < i \<longrightarrow> Array.get h (a::nat array) ! i = Array.get h' a ! i" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 268 | proof - | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 269 | from assms part_outer_remains part_returns_index_in_bounds show ?thesis | 
| 28145 | 270 | unfolding partition.simps swap_def | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 271 | by (elim effect_bindE effect_returnE effect_nthE effect_ifE effect_updE) fastforce | 
| 27656 
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 partition_returns_index_in_bounds: | 
| 40671 | 275 | assumes effect: "effect (partition a l r) h h' rs" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 276 | assumes "l < r" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 277 | shows "l \<le> rs \<and> rs \<le> r" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 278 | proof - | 
| 40671 | 279 | from effect obtain middle h'' p where part: "effect (part1 a l (r - 1) p) h h'' middle" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 280 | and rs_equals: "rs = (if Array.get h'' a ! middle \<le> Array.get h a ! r then middle + 1 | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 281 | else middle)" | 
| 28145 | 282 | unfolding partition.simps | 
| 40671 | 283 | by (elim effect_bindE effect_returnE effect_nthE effect_ifE effect_updE) simp | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 284 | from `l < r` have "l \<le> r - 1" by arith | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 285 | from part_returns_index_in_bounds[OF part this] rs_equals `l < r` show ?thesis by auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 286 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 287 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 288 | lemma partition_partitions: | 
| 40671 | 289 | assumes effect: "effect (partition a l r) h h' rs" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 290 | assumes "l < r" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 291 | shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> Array.get h' (a::nat array) ! i \<le> Array.get h' a ! rs) \<and> | 
| 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 292 | (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> Array.get h' a ! rs \<le> Array.get h' a ! i)" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 293 | proof - | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 294 | let ?pivot = "Array.get h a ! r" | 
| 40671 | 295 | from effect obtain middle h1 where part: "effect (part1 a l (r - 1) ?pivot) h h1 middle" | 
| 296 | and swap: "effect (swap a rs r) h1 h' ()" | |
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 297 | and rs_equals: "rs = (if Array.get h1 a ! middle \<le> ?pivot then middle + 1 | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 298 | else middle)" | 
| 28145 | 299 | unfolding partition.simps | 
| 40671 | 300 | by (elim effect_bindE effect_returnE effect_nthE effect_ifE effect_updE) simp | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 301 | from swap have h'_def: "h' = Array.update a r (Array.get h1 a ! rs) | 
| 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 302 | (Array.update a rs (Array.get h1 a ! r) h1)" | 
| 28145 | 303 | unfolding swap_def | 
| 40671 | 304 | by (elim effect_bindE effect_returnE effect_nthE effect_updE) simp | 
| 37802 | 305 | from swap have in_bounds: "r < Array.length h1 a \<and> rs < Array.length h1 a" | 
| 28145 | 306 | unfolding swap_def | 
| 40671 | 307 | by (elim effect_bindE effect_returnE effect_nthE effect_updE) simp | 
| 37802 | 308 | from swap have swap_length_remains: "Array.length h1 a = Array.length h' a" | 
| 40671 | 309 | unfolding swap_def by (elim effect_bindE effect_returnE effect_nthE effect_updE) auto | 
| 37771 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37758diff
changeset | 310 | from `l < r` have "l \<le> r - 1" by simp | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 311 | note middle_in_bounds = part_returns_index_in_bounds[OF part this] | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 312 | from part_outer_remains[OF part] `l < r` | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 313 | have "Array.get h a ! r = Array.get h1 a ! r" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 314 | by fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 315 | with swap | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 316 | have right_remains: "Array.get h a ! r = Array.get h' a ! rs" | 
| 28145 | 317 | unfolding swap_def | 
| 40671 | 318 | by (auto simp add: Array.length_def elim!: effect_bindE effect_returnE effect_nthE effect_updE) (cases "r = rs", auto) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 319 | from part_partitions [OF part] | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 320 | show ?thesis | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 321 | proof (cases "Array.get h1 a ! middle \<le> ?pivot") | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 322 | case True | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 323 | with rs_equals have rs_equals: "rs = middle + 1" by simp | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 324 |     { 
 | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 325 | fix i | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 326 | assume i_is_left: "l \<le> i \<and> i < rs" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 327 | with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r` | 
| 37802 | 328 | have i_props: "i < Array.length h' a" "i \<noteq> r" "i \<noteq> rs" by auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 329 | from i_is_left rs_equals have "l \<le> i \<and> i < middle \<or> i = middle" by arith | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 330 | with part_partitions[OF part] right_remains True | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 331 | have "Array.get h1 a ! i \<le> Array.get h' a ! rs" by fastforce | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 332 | with i_props h'_def in_bounds have "Array.get h' a ! i \<le> Array.get h' a ! rs" | 
| 37798 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 333 | unfolding Array.update_def Array.length_def by simp | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 334 | } | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 335 | moreover | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 336 |     {
 | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 337 | fix i | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 338 | assume "rs < i \<and> i \<le> r" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 339 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 340 | hence "(rs < i \<and> i \<le> r - 1) \<or> (rs < i \<and> i = r)" by arith | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 341 | hence "Array.get h' a ! rs \<le> Array.get h' a ! i" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 342 | proof | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 343 | assume i_is: "rs < i \<and> i \<le> r - 1" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 344 | with swap_length_remains in_bounds middle_in_bounds rs_equals | 
| 37802 | 345 | have i_props: "i < Array.length h' a" "i \<noteq> r" "i \<noteq> rs" by auto | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 346 | from part_partitions[OF part] rs_equals right_remains i_is | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 347 | have "Array.get h' a ! rs \<le> Array.get h1 a ! i" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 348 | by fastforce | 
| 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 349 | with i_props h'_def show ?thesis by fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 350 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 351 | assume i_is: "rs < i \<and> i = r" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 352 | with rs_equals have "Suc middle \<noteq> r" by arith | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 353 | with middle_in_bounds `l < r` have "Suc middle \<le> r - 1" by arith | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 354 | with part_partitions[OF part] right_remains | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 355 | have "Array.get h' a ! rs \<le> Array.get h1 a ! (Suc middle)" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 356 | by fastforce | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 357 | with i_is True rs_equals right_remains h'_def | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 358 | show ?thesis using in_bounds | 
| 37798 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 359 | unfolding Array.update_def Array.length_def | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 360 | by auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 361 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 362 | } | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 363 | ultimately show ?thesis by auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 364 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 365 | case False | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 366 | with rs_equals have rs_equals: "middle = rs" by simp | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 367 |     { 
 | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 368 | fix i | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 369 | assume i_is_left: "l \<le> i \<and> i < rs" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 370 | with swap_length_remains in_bounds middle_in_bounds rs_equals | 
| 37802 | 371 | have i_props: "i < Array.length h' a" "i \<noteq> r" "i \<noteq> rs" by auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 372 | from part_partitions[OF part] rs_equals right_remains i_is_left | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 373 | have "Array.get h1 a ! i \<le> Array.get h' a ! rs" by fastforce | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 374 | with i_props h'_def have "Array.get h' a ! i \<le> Array.get h' a ! rs" | 
| 37798 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 375 | unfolding Array.update_def by simp | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 376 | } | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 377 | moreover | 
| 
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 | fix i | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 380 | assume "rs < i \<and> i \<le> r" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 381 | hence "(rs < i \<and> i \<le> r - 1) \<or> i = r" by arith | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 382 | hence "Array.get h' a ! rs \<le> Array.get h' a ! i" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 383 | proof | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 384 | assume i_is: "rs < i \<and> i \<le> r - 1" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 385 | with swap_length_remains in_bounds middle_in_bounds rs_equals | 
| 37802 | 386 | have i_props: "i < Array.length h' a" "i \<noteq> r" "i \<noteq> rs" by auto | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 387 | from part_partitions[OF part] rs_equals right_remains i_is | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 388 | have "Array.get h' a ! rs \<le> Array.get h1 a ! i" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 389 | by fastforce | 
| 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 390 | with i_props h'_def show ?thesis by fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 391 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 392 | assume i_is: "i = r" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 393 | from i_is False rs_equals right_remains h'_def | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 394 | show ?thesis using in_bounds | 
| 37798 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 395 | unfolding Array.update_def Array.length_def | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 396 | by auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 397 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 398 | } | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 399 | ultimately | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 400 | show ?thesis by auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 401 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 402 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 403 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 404 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 405 | function quicksort :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 406 | where | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 407 | "quicksort arr left right = | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 408 | (if (right > left) then | 
| 37792 | 409 |         do {
 | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 410 | pivotNewIndex \<leftarrow> partition arr left right; | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 411 | pivotNewIndex \<leftarrow> assert (\<lambda>x. left \<le> x \<and> x \<le> right) pivotNewIndex; | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 412 | quicksort arr left (pivotNewIndex - 1); | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 413 | quicksort arr (pivotNewIndex + 1) right | 
| 37792 | 414 | } | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 415 | else return ())" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 416 | by pat_completeness auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 417 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 418 | (* For termination, we must show that the pivotNewIndex is between left and right *) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 419 | termination | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 420 | by (relation "measure (\<lambda>(a, l, r). (r - l))") auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 421 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 422 | declare quicksort.simps[simp del] | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 423 | |
| 
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 quicksort_permutes: | 
| 40671 | 426 | assumes "effect (quicksort a l r) h h' rs" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 427 | shows "multiset_of (Array.get h' a) | 
| 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 428 | = multiset_of (Array.get h a)" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 429 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 430 | proof (induct a l r arbitrary: h h' rs rule: quicksort.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 431 | case (1 a l r h h' rs) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 432 | with partition_permutes show ?case | 
| 28145 | 433 | unfolding quicksort.simps [of a l r] | 
| 40671 | 434 | by (elim effect_ifE effect_bindE effect_assertE effect_returnE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 435 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 436 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 437 | lemma length_remains: | 
| 40671 | 438 | assumes "effect (quicksort a l r) h h' rs" | 
| 37802 | 439 | shows "Array.length h a = Array.length h' a" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 440 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 441 | proof (induct a l r arbitrary: h h' rs rule: quicksort.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 442 | case (1 a l r h h' rs) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 443 | with partition_length_remains show ?case | 
| 28145 | 444 | unfolding quicksort.simps [of a l r] | 
| 40671 | 445 | by (elim effect_ifE effect_bindE effect_assertE effect_returnE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 446 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 447 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 448 | lemma quicksort_outer_remains: | 
| 40671 | 449 | assumes "effect (quicksort a l r) h h' rs" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 450 | shows "\<forall>i. i < l \<or> r < i \<longrightarrow> Array.get h (a::nat array) ! i = Array.get h' a ! i" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 451 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 452 | proof (induct a l r arbitrary: h h' rs rule: quicksort.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 453 | case (1 a l r h h' rs) | 
| 40671 | 454 | note cr = `effect (quicksort a l r) h h' rs` | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 455 | thus ?case | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 456 | proof (cases "r > l") | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 457 | case False | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 458 | with cr have "h' = h" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 459 | unfolding quicksort.simps [of a l r] | 
| 40671 | 460 | by (elim effect_ifE effect_returnE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 461 | thus ?thesis by simp | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 462 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 463 | case True | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 464 |    { 
 | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 465 | fix h1 h2 p ret1 ret2 i | 
| 40671 | 466 | assume part: "effect (partition a l r) h h1 p" | 
| 467 | assume qs1: "effect (quicksort a l (p - 1)) h1 h2 ret1" | |
| 468 | assume qs2: "effect (quicksort a (p + 1) r) h2 h' ret2" | |
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 469 | assume pivot: "l \<le> p \<and> p \<le> r" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 470 | assume i_outer: "i < l \<or> r < i" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 471 | from partition_outer_remains [OF part True] i_outer | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 472 | have "Array.get h a !i = Array.get h1 a ! i" by fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 473 | moreover | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 474 | with 1(1) [OF True pivot qs1] pivot i_outer | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 475 | have "Array.get h1 a ! i = Array.get h2 a ! i" by auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 476 | moreover | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 477 | with qs2 1(2) [of p h2 h' ret2] True pivot i_outer | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 478 | have "Array.get h2 a ! i = Array.get h' a ! i" by auto | 
| 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 479 | ultimately have "Array.get h a ! i= Array.get h' a ! i" by simp | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 480 | } | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 481 | with cr show ?thesis | 
| 28145 | 482 | unfolding quicksort.simps [of a l r] | 
| 40671 | 483 | by (elim effect_ifE effect_bindE effect_assertE effect_returnE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 484 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 485 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 486 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 487 | lemma quicksort_is_skip: | 
| 40671 | 488 | assumes "effect (quicksort a l r) h h' rs" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 489 | shows "r \<le> l \<longrightarrow> h = h'" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 490 | using assms | 
| 28145 | 491 | unfolding quicksort.simps [of a l r] | 
| 40671 | 492 | by (elim effect_ifE effect_returnE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 493 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 494 | lemma quicksort_sorts: | 
| 40671 | 495 | assumes "effect (quicksort a l r) h h' rs" | 
| 37802 | 496 | assumes l_r_length: "l < Array.length h a" "r < Array.length h a" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 497 | shows "sorted (subarray l (r + 1) a h')" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 498 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 499 | proof (induct a l r arbitrary: h h' rs rule: quicksort.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 500 | case (1 a l r h h' rs) | 
| 40671 | 501 | note cr = `effect (quicksort a l r) h h' rs` | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 502 | thus ?case | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 503 | proof (cases "r > l") | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 504 | case False | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 505 | hence "l \<ge> r + 1 \<or> l = r" by arith | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 506 | with length_remains[OF cr] 1(5) show ?thesis | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 507 | by (auto simp add: subarray_Nil subarray_single) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 508 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 509 | case True | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 510 |     { 
 | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 511 | fix h1 h2 p | 
| 40671 | 512 | assume part: "effect (partition a l r) h h1 p" | 
| 513 | assume qs1: "effect (quicksort a l (p - 1)) h1 h2 ()" | |
| 514 | assume qs2: "effect (quicksort a (p + 1) r) h2 h' ()" | |
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 515 | from partition_returns_index_in_bounds [OF part True] | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 516 | have pivot: "l\<le> p \<and> p \<le> r" . | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 517 | note length_remains = length_remains[OF qs2] length_remains[OF qs1] partition_length_remains[OF part] | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 518 | from quicksort_outer_remains [OF qs2] quicksort_outer_remains [OF qs1] pivot quicksort_is_skip[OF qs1] | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 519 | have pivot_unchanged: "Array.get h1 a ! p = Array.get h' a ! p" by (cases p, auto) | 
| 28013 | 520 | (*-- First of all, by induction hypothesis both sublists are sorted. *) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 521 | from 1(1)[OF True pivot qs1] length_remains pivot 1(5) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 522 | have IH1: "sorted (subarray l p a h2)" by (cases p, auto simp add: subarray_Nil) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 523 | from quicksort_outer_remains [OF qs2] length_remains | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 524 | have left_subarray_remains: "subarray l p a h2 = subarray l p a h'" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 525 | by (simp add: subarray_eq_samelength_iff) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 526 | with IH1 have IH1': "sorted (subarray l p a h')" by simp | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 527 | from 1(2)[OF True pivot qs2] pivot 1(5) length_remains | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 528 | have IH2: "sorted (subarray (p + 1) (r + 1) a h')" | 
| 28013 | 529 | by (cases "Suc p \<le> r", auto simp add: subarray_Nil) | 
| 530 | (* -- Secondly, both sublists remain partitioned. *) | |
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 531 | from partition_partitions[OF part True] | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 532 | have part_conds1: "\<forall>j. j \<in> set (subarray l p a h1) \<longrightarrow> j \<le> Array.get h1 a ! p " | 
| 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 533 | and part_conds2: "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h1) \<longrightarrow> Array.get h1 a ! p \<le> j" | 
| 28013 | 534 | by (auto simp add: all_in_set_subarray_conv) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 535 | from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 536 | length_remains 1(5) pivot multiset_of_sublist [of l p "Array.get h1 a" "Array.get h2 a"] | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 537 | have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)" | 
| 37719 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 haftmann parents: 
36098diff
changeset | 538 | unfolding Array.length_def subarray_def by (cases p, auto) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 539 | with left_subarray_remains part_conds1 pivot_unchanged | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 540 | have part_conds2': "\<forall>j. j \<in> set (subarray l p a h') \<longrightarrow> j \<le> Array.get h' a ! p" | 
| 28013 | 541 | by (simp, subst set_of_multiset_of[symmetric], simp) | 
| 542 | (* -- These steps are the analogous for the right sublist \<dots> *) | |
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 543 | from quicksort_outer_remains [OF qs1] length_remains | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 544 | have right_subarray_remains: "subarray (p + 1) (r + 1) a h1 = subarray (p + 1) (r + 1) a h2" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 545 | by (auto simp add: subarray_eq_samelength_iff) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 546 | from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 547 | length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "Array.get h2 a" "Array.get h' a"] | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 548 | have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)" | 
| 37719 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 haftmann parents: 
36098diff
changeset | 549 | unfolding Array.length_def subarray_def by auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 550 | with right_subarray_remains part_conds2 pivot_unchanged | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 551 | have part_conds1': "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h') \<longrightarrow> Array.get h' a ! p \<le> j" | 
| 28013 | 552 | by (simp, subst set_of_multiset_of[symmetric], simp) | 
| 553 | (* -- Thirdly and finally, we show that the array is sorted | |
| 554 | following from the facts above. *) | |
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 555 | from True pivot 1(5) length_remains have "subarray l (r + 1) a h' = subarray l p a h' @ [Array.get h' a ! p] @ subarray (p + 1) (r + 1) a h'" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 556 | by (simp add: subarray_nth_array_Cons, cases "l < p") (auto simp add: subarray_append subarray_Nil) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 557 | with IH1' IH2 part_conds1' part_conds2' pivot have ?thesis | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 558 | unfolding subarray_def | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31887diff
changeset | 559 | apply (auto simp add: sorted_append sorted_Cons all_in_set_sublist'_conv) | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 560 | by (auto simp add: set_sublist' dest: le_trans [of _ "Array.get h' a ! p"]) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 561 | } | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 562 | with True cr show ?thesis | 
| 28145 | 563 | unfolding quicksort.simps [of a l r] | 
| 40671 | 564 | by (elim effect_ifE effect_returnE effect_bindE effect_assertE) auto | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 565 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 566 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 567 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 568 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 569 | lemma quicksort_is_sort: | 
| 40671 | 570 | assumes effect: "effect (quicksort a 0 (Array.length h a - 1)) h h' rs" | 
| 37806 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 571 | shows "Array.get h' a = sort (Array.get h a)" | 
| 
a7679be14442
qualified names for (really) all array operations
 haftmann parents: 
37805diff
changeset | 572 | proof (cases "Array.get h a = []") | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 573 | case True | 
| 40671 | 574 | with quicksort_is_skip[OF effect] show ?thesis | 
| 37719 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 haftmann parents: 
36098diff
changeset | 575 | unfolding Array.length_def by simp | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 576 | next | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 577 | case False | 
| 40671 | 578 | from quicksort_sorts [OF effect] False have "sorted (sublist' 0 (List.length (Array.get h a)) (Array.get h' a))" | 
| 37719 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 haftmann parents: 
36098diff
changeset | 579 | unfolding Array.length_def subarray_def by auto | 
| 40671 | 580 | with length_remains[OF effect] have "sorted (Array.get h' a)" | 
| 37719 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 haftmann parents: 
36098diff
changeset | 581 | unfolding Array.length_def by simp | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41413diff
changeset | 582 | with quicksort_permutes [OF effect] properties_for_sort show ?thesis by fastforce | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 583 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 584 | |
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 585 | subsection {* No Errors in quicksort *}
 | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 586 | text {* We have proved that quicksort sorts (if no exceptions occur).
 | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 587 | We will now show that exceptions do not occur. *} | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 588 | |
| 37758 | 589 | lemma success_part1I: | 
| 37802 | 590 | assumes "l < Array.length h a" "r < Array.length h a" | 
| 37758 | 591 | shows "success (part1 a l r p) h" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 592 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 593 | proof (induct a l r p arbitrary: h rule: part1.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 594 | case (1 a l r p) | 
| 37771 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37758diff
changeset | 595 | thus ?case unfolding part1.simps [of a l r] | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37758diff
changeset | 596 | apply (auto intro!: success_intros del: success_ifI simp add: not_le) | 
| 40671 | 597 | apply (auto intro!: effect_intros effect_swapI) | 
| 37771 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37758diff
changeset | 598 | done | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 599 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 600 | |
| 37758 | 601 | lemma success_bindI' [success_intros]: (*FIXME move*) | 
| 602 | assumes "success f h" | |
| 40671 | 603 | assumes "\<And>h' r. effect f h h' r \<Longrightarrow> success (g r) h'" | 
| 37758 | 604 | shows "success (f \<guillemotright>= g) h" | 
| 40671 | 605 | using assms(1) proof (rule success_effectE) | 
| 37771 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37758diff
changeset | 606 | fix h' r | 
| 40671 | 607 | assume "effect f h h' r" | 
| 37771 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37758diff
changeset | 608 | moreover with assms(2) have "success (g r) h'" . | 
| 40671 | 609 | ultimately show "success (f \<guillemotright>= g) h" by (rule success_bind_effectI) | 
| 37771 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37758diff
changeset | 610 | qed | 
| 37758 | 611 | |
| 612 | lemma success_partitionI: | |
| 37802 | 613 | assumes "l < r" "l < Array.length h a" "r < Array.length h a" | 
| 37758 | 614 | shows "success (partition a l r) h" | 
| 615 | using assms unfolding partition.simps swap_def | |
| 40671 | 616 | apply (auto intro!: success_bindI' success_ifI success_returnI success_nthI success_updI success_part1I elim!: effect_bindE effect_updE effect_nthE effect_returnE simp add:) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 617 | apply (frule part_length_remains) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 618 | apply (frule part_returns_index_in_bounds) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 619 | apply auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 620 | apply (frule part_length_remains) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 621 | apply (frule part_returns_index_in_bounds) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 622 | apply auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 623 | apply (frule part_length_remains) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 624 | apply auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 625 | done | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 626 | |
| 37758 | 627 | lemma success_quicksortI: | 
| 37802 | 628 | assumes "l < Array.length h a" "r < Array.length h a" | 
| 37758 | 629 | shows "success (quicksort a l r) h" | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 630 | using assms | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 631 | proof (induct a l r arbitrary: h rule: quicksort.induct) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 632 | case (1 a l ri h) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 633 | thus ?case | 
| 28145 | 634 | unfolding quicksort.simps [of a l ri] | 
| 37758 | 635 | apply (auto intro!: success_ifI success_bindI' success_returnI success_nthI success_updI success_assertI success_partitionI) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 636 | apply (frule partition_returns_index_in_bounds) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 637 | apply auto | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 638 | apply (frule partition_returns_index_in_bounds) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 639 | apply auto | 
| 40671 | 640 | apply (auto elim!: effect_assertE dest!: partition_length_remains length_remains) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 641 | apply (subgoal_tac "Suc r \<le> ri \<or> r = ri") | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 642 | apply (erule disjE) | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 643 | apply auto | 
| 28145 | 644 | unfolding quicksort.simps [of a "Suc ri" ri] | 
| 37758 | 645 | apply (auto intro!: success_ifI success_returnI) | 
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 646 | done | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 647 | qed | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: diff
changeset | 648 | |
| 27674 | 649 | |
| 650 | subsection {* Example *}
 | |
| 651 | ||
| 37792 | 652 | definition "qsort a = do {
 | 
| 37798 
0b0570445a2a
proper merge of operation changes and generic do-syntax
 haftmann parents: 
37797diff
changeset | 653 | k \<leftarrow> Array.len a; | 
| 27674 | 654 | quicksort a 0 (k - 1); | 
| 655 | return a | |
| 37792 | 656 | }" | 
| 27674 | 657 | |
| 35041 
6eb917794a5c
avoid upto in generated code (is infix operator in library.ML)
 haftmann parents: 
32960diff
changeset | 658 | code_reserved SML upto | 
| 
6eb917794a5c
avoid upto in generated code (is infix operator in library.ML)
 haftmann parents: 
32960diff
changeset | 659 | |
| 27674 | 660 | ML {* @{code qsort} (Array.fromList [42, 2, 3, 5, 0, 1705, 8, 3, 15]) () *}
 | 
| 661 | ||
| 38058 | 662 | export_code qsort checking SML SML_imp OCaml? OCaml_imp? Haskell? Scala? Scala_imp? | 
| 27674 | 663 | |
| 37845 
b70d7a347964
first roughly working version of Imperative HOL for Scala
 haftmann parents: 
37826diff
changeset | 664 | end |