src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy
author haftmann
Mon, 23 Mar 2009 19:01:17 +0100
changeset 30689 b14b2cc4e25e
parent 29793 src/HOL/ex/ImperativeQuicksort.thy@86cac1fab613
child 31873 00878e406bf5
permissions -rw-r--r--
moved Imperative_HOL examples to Imperative_HOL/ex
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30689
b14b2cc4e25e moved Imperative_HOL examples to Imperative_HOL/ex
haftmann
parents: 29793
diff changeset
     1
(* Author: Lukas Bulwahn, TU Muenchen *)
b14b2cc4e25e moved Imperative_HOL examples to Imperative_HOL/ex
haftmann
parents: 29793
diff changeset
     2
b14b2cc4e25e moved Imperative_HOL examples to Imperative_HOL/ex
haftmann
parents: 29793
diff changeset
     3
theory Imperative_Quicksort
29399
ebcd69a00872 split of Imperative_HOL theories from HOL-Library
haftmann
parents: 28145
diff changeset
     4
imports "~~/src/HOL/Imperative_HOL/Imperative_HOL" Subarray Multiset Efficient_Nat
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     5
begin
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     6
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     7
text {* 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
     8
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     9
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
    10
where
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    11
  "swap arr i j = (
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    12
     do
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    13
       x \<leftarrow> nth arr i;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    14
       y \<leftarrow> nth arr j;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    15
       upd i y arr;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    16
       upd j x arr;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    17
       return ()
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    18
     done)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    19
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    20
lemma swap_permutes:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    21
  assumes "crel (swap a i j) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    22
  shows "multiset_of (get_array a h') 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    23
  = multiset_of (get_array a h)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    24
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
    25
  unfolding swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    26
  by (auto simp add: Heap.length_def multiset_of_swap dest: sym [of _ "h'"] elim!: crelE crel_nth crel_return crel_upd)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    27
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    28
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
    29
where
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    30
  "part1 a left right p = (
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    31
     if (right \<le> left) then return right
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    32
     else (do
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    33
       v \<leftarrow> nth a left;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    34
       (if (v \<le> p) then (part1 a (left + 1) right p)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    35
                    else (do swap a left right;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    36
  part1 a left (right - 1) p done))
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    37
     done))"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    38
by pat_completeness auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    39
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    40
termination
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    41
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
    42
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    43
declare part1.simps[simp del]
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    44
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    45
lemma part_permutes:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    46
  assumes "crel (part1 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
    47
  shows "multiset_of (get_array a h') 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    48
  = multiset_of (get_array a h)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    49
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    50
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
    51
  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
    52
  thus ?case
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
    53
    unfolding part1.simps [of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    54
    by (elim crelE crel_if crel_return crel_nth) (auto simp add: swap_permutes)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    55
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    56
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    57
lemma part_returns_index_in_bounds:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    58
  assumes "crel (part1 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
    59
  assumes "l \<le> r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    60
  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
    61
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    62
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
    63
  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
    64
  note cr = `crel (part1 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
    65
  show ?case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    66
  proof (cases "r \<le> l")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    67
    case True (* Terminating case *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    68
    with cr `l \<le> r` show ?thesis
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
    69
      unfolding part1.simps[of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    70
      by (elim crelE crel_if crel_return crel_nth) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    71
  next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    72
    case False (* recursive case *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    73
    note rec_condition = this
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    74
    let ?v = "get_array a h ! l"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    75
    show ?thesis
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    76
    proof (cases "?v \<le> p")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    77
      case True
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    78
      with cr False
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    79
      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
    80
        unfolding part1.simps[of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    81
        by (elim crelE crel_nth crel_if crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    82
      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
    83
      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
    84
      show ?thesis by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    85
    next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    86
      case False
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    87
      with rec_condition cr
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    88
      obtain h1 where swp: "crel (swap a l r) h h1 ()"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    89
        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
    90
        unfolding part1.simps[of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    91
        by (elim crelE crel_nth crel_if crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    92
      from rec_condition 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
    93
      from 1(2) [OF rec_condition False rec2 `l \<le> r - 1`] show ?thesis by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    94
    qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    95
  qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    96
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    97
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    98
lemma part_length_remains:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    99
  assumes "crel (part1 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
   100
  shows "Heap.length a h = Heap.length a h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   101
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   102
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
   103
  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
   104
  note cr = `crel (part1 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
   105
  
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   106
  show ?case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   107
  proof (cases "r \<le> l")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   108
    case True (* Terminating case *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   109
    with cr show ?thesis
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   110
      unfolding part1.simps[of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   111
      by (elim crelE crel_if crel_return crel_nth) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   112
  next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   113
    case False (* recursive case *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   114
    with cr 1 show ?thesis
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   115
      unfolding part1.simps [of a l r p] swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   116
      by (auto elim!: crelE crel_if crel_nth crel_return crel_upd) fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   117
  qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   118
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   119
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   120
lemma part_outer_remains:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   121
  assumes "crel (part1 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
   122
  shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   123
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   124
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
   125
  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
   126
  note cr = `crel (part1 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
   127
  
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   128
  show ?case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   129
  proof (cases "r \<le> l")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   130
    case True (* Terminating case *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   131
    with cr show ?thesis
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   132
      unfolding part1.simps[of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   133
      by (elim crelE crel_if crel_return crel_nth) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   134
  next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   135
    case False (* recursive case *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   136
    note rec_condition = this
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   137
    let ?v = "get_array a h ! l"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   138
    show ?thesis
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   139
    proof (cases "?v \<le> p")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   140
      case True
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   141
      with cr False
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   142
      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   143
        unfolding part1.simps[of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   144
        by (elim crelE crel_nth crel_if crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   145
      from 1(1)[OF rec_condition True rec1]
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   146
      show ?thesis by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   147
    next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   148
      case False
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   149
      with rec_condition cr
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   150
      obtain h1 where swp: "crel (swap a l r) h h1 ()"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   151
        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   152
        unfolding part1.simps[of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   153
        by (elim crelE crel_nth crel_if crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   154
      from swp rec_condition have
28013
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   155
        "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array a h ! i = get_array a h1 ! i"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   156
	unfolding swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   157
	by (elim crelE crel_nth crel_upd crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   158
      with 1(2) [OF rec_condition False rec2] show ?thesis by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   159
    qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   160
  qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   161
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   162
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   163
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   164
lemma part_partitions:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   165
  assumes "crel (part1 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
   166
  shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> p)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   167
  \<and> (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! i \<ge> p)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   168
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   169
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
   170
  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
   171
  note cr = `crel (part1 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
   172
  
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   173
  show ?case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   174
  proof (cases "r \<le> l")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   175
    case True (* Terminating case *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   176
    with cr have "rs = r"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   177
      unfolding part1.simps[of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   178
      by (elim crelE crel_if crel_return crel_nth) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   179
    with True
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   180
    show ?thesis by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   181
  next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   182
    case False (* recursive case *)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   183
    note lr = this
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   184
    let ?v = "get_array a h ! l"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   185
    show ?thesis
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   186
    proof (cases "?v \<le> p")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   187
      case True
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   188
      with lr cr
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   189
      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   190
        unfolding part1.simps[of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   191
        by (elim crelE crel_nth crel_if crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   192
      from True part_outer_remains[OF rec1] have a_l: "get_array a h' ! l \<le> p"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   193
	by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   194
      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
   195
      with 1(1)[OF False True rec1] a_l show ?thesis
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   196
	by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   197
    next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   198
      case False
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   199
      with lr cr
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   200
      obtain h1 where swp: "crel (swap a l r) h h1 ()"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   201
        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   202
        unfolding part1.simps[of a l r p]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   203
        by (elim crelE crel_nth crel_if crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   204
      from swp False have "get_array a h1 ! r \<ge> p"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   205
	unfolding swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   206
	by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   207
      with part_outer_remains [OF rec2] lr have a_r: "get_array a h' ! r \<ge> p"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   208
	by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   209
      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
   210
      with 1(2)[OF lr False rec2] a_r show ?thesis
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   211
	by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   212
    qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   213
  qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   214
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   215
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   216
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   217
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
   218
where
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   219
  "partition a left right = (do
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   220
     pivot \<leftarrow> nth a right;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   221
     middle \<leftarrow> part1 a left (right - 1) pivot;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   222
     v \<leftarrow> nth a middle;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   223
     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
   224
     swap a m right;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   225
     return m
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   226
   done)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   227
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   228
declare partition.simps[simp del]
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   229
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   230
lemma partition_permutes:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   231
  assumes "crel (partition a l r) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   232
  shows "multiset_of (get_array a h') 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   233
  = multiset_of (get_array a h)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   234
proof -
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   235
    from assms part_permutes swap_permutes show ?thesis
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   236
      unfolding partition.simps
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   237
      by (elim crelE crel_return crel_nth crel_if crel_upd) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   238
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   239
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   240
lemma partition_length_remains:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   241
  assumes "crel (partition a l r) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   242
  shows "Heap.length a h = Heap.length a h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   243
proof -
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   244
  from assms part_length_remains show ?thesis
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   245
    unfolding partition.simps swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   246
    by (elim crelE crel_return crel_nth crel_if crel_upd) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   247
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   248
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   249
lemma partition_outer_remains:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   250
  assumes "crel (partition a l r) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   251
  assumes "l < r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   252
  shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   253
proof -
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   254
  from assms part_outer_remains part_returns_index_in_bounds show ?thesis
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   255
    unfolding partition.simps swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   256
    by (elim crelE crel_return crel_nth crel_if crel_upd) fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   257
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   258
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   259
lemma partition_returns_index_in_bounds:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   260
  assumes crel: "crel (partition a l r) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   261
  assumes "l < r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   262
  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
   263
proof -
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   264
  from crel obtain middle h'' p where part: "crel (part1 a l (r - 1) p) h h'' middle"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   265
    and rs_equals: "rs = (if get_array a h'' ! middle \<le> get_array a h ! r then middle + 1
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   266
         else middle)"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   267
    unfolding partition.simps
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   268
    by (elim crelE crel_return crel_nth crel_if crel_upd) simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   269
  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
   270
  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
   271
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   272
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   273
lemma partition_partitions:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   274
  assumes crel: "crel (partition a l r) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   275
  assumes "l < r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   276
  shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> get_array a h' ! rs) \<and>
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   277
  (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! rs \<le> get_array a h' ! i)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   278
proof -
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   279
  let ?pivot = "get_array a h ! r" 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   280
  from crel obtain middle h1 where part: "crel (part1 a l (r - 1) ?pivot) h h1 middle"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   281
    and swap: "crel (swap a rs r) h1 h' ()"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   282
    and rs_equals: "rs = (if get_array a h1 ! middle \<le> ?pivot then middle + 1
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   283
         else middle)"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   284
    unfolding partition.simps
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   285
    by (elim crelE crel_return crel_nth crel_if crel_upd) simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   286
  from swap have h'_def: "h' = Heap.upd a r (get_array a h1 ! rs)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   287
    (Heap.upd a rs (get_array a h1 ! r) h1)"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   288
    unfolding swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   289
    by (elim crelE crel_return crel_nth crel_upd) simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   290
  from swap have in_bounds: "r < Heap.length a h1 \<and> rs < Heap.length a h1"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   291
    unfolding swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   292
    by (elim crelE crel_return crel_nth crel_upd) simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   293
  from swap have swap_length_remains: "Heap.length a h1 = Heap.length a h'"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   294
    unfolding swap_def by (elim crelE crel_return crel_nth crel_upd) auto
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   295
  from `l < r` have "l \<le> r - 1" by simp 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   296
  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
   297
  from part_outer_remains[OF part] `l < r`
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   298
  have "get_array a h ! r = get_array a h1 ! r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   299
    by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   300
  with swap
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   301
  have right_remains: "get_array a h ! r = get_array a h' ! rs"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   302
    unfolding swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   303
    by (auto simp add: Heap.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   304
  from part_partitions [OF part]
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   305
  show ?thesis
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   306
  proof (cases "get_array a h1 ! middle \<le> ?pivot")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   307
    case True
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   308
    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
   309
    { 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   310
      fix i
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   311
      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
   312
      with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r`
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   313
      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   314
      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
   315
      with part_partitions[OF part] right_remains True
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   316
      have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   317
      with i_props h'_def in_bounds have "get_array a h' ! i \<le> get_array a h' ! rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   318
	unfolding Heap.upd_def Heap.length_def by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   319
    }
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   320
    moreover
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   321
    {
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   322
      fix i
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   323
      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
   324
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   325
      hence "(rs < i \<and> i \<le> r - 1) \<or> (rs < i \<and> i = r)" by arith
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   326
      hence "get_array a h' ! rs \<le> get_array a h' ! i"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   327
      proof
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   328
	assume i_is: "rs < i \<and> i \<le> r - 1"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   329
	with swap_length_remains in_bounds middle_in_bounds rs_equals
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   330
	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   331
	from part_partitions[OF part] rs_equals right_remains i_is
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   332
	have "get_array a h' ! rs \<le> get_array a h1 ! i"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   333
	  by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   334
	with i_props h'_def show ?thesis by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   335
      next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   336
	assume i_is: "rs < i \<and> i = r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   337
	with rs_equals have "Suc middle \<noteq> r" by arith
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   338
	with middle_in_bounds `l < r` have "Suc middle \<le> r - 1" by arith
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   339
	with part_partitions[OF part] right_remains 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   340
	have "get_array a h' ! rs \<le> get_array a h1 ! (Suc middle)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   341
	  by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   342
	with i_is True rs_equals right_remains h'_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   343
	show ?thesis using in_bounds
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   344
	  unfolding Heap.upd_def Heap.length_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   345
	  by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   346
      qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   347
    }
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   348
    ultimately show ?thesis by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   349
  next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   350
    case False
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   351
    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
   352
    { 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   353
      fix i
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   354
      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
   355
      with swap_length_remains in_bounds middle_in_bounds rs_equals
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   356
      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   357
      from part_partitions[OF part] rs_equals right_remains i_is_left
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   358
      have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   359
      with i_props h'_def have "get_array a h' ! i \<le> get_array a h' ! rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   360
	unfolding Heap.upd_def by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   361
    }
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   362
    moreover
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   363
    {
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   364
      fix i
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   365
      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
   366
      hence "(rs < i \<and> i \<le> r - 1) \<or> i = r" by arith
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   367
      hence "get_array a h' ! rs \<le> get_array a h' ! i"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   368
      proof
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   369
	assume i_is: "rs < i \<and> i \<le> r - 1"
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
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   371
	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
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
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   373
	have "get_array a h' ! rs \<le> get_array a h1 ! i"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   374
	  by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   375
	with i_props h'_def show ?thesis by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   376
      next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   377
	assume i_is: "i = r"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   378
	from i_is False rs_equals right_remains h'_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   379
	show ?thesis using in_bounds
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   380
	  unfolding Heap.upd_def Heap.length_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   381
	  by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   382
      qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   383
    }
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   384
    ultimately
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   385
    show ?thesis by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   386
  qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   387
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   388
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   389
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   390
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
   391
where
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   392
  "quicksort arr left right =
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   393
     (if (right > left)  then
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   394
        do
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   395
          pivotNewIndex \<leftarrow> partition arr left right;
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   396
          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
   397
          quicksort arr left (pivotNewIndex - 1);
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   398
          quicksort arr (pivotNewIndex + 1) right
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   399
        done
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   400
     else return ())"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   401
by pat_completeness auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   402
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   403
(* 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
   404
termination
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   405
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
   406
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   407
declare quicksort.simps[simp del]
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   408
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   409
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   410
lemma quicksort_permutes:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   411
  assumes "crel (quicksort a l r) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   412
  shows "multiset_of (get_array a h') 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   413
  = multiset_of (get_array a h)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   414
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   415
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
   416
  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
   417
  with partition_permutes show ?case
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   418
    unfolding quicksort.simps [of a l r]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   419
    by (elim crel_if crelE crel_assert crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   420
qed
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
lemma length_remains:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   423
  assumes "crel (quicksort a l r) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   424
  shows "Heap.length a h = Heap.length a h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   425
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   426
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
   427
  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
   428
  with partition_length_remains show ?case
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   429
    unfolding quicksort.simps [of a l r]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   430
    by (elim crel_if crelE crel_assert crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   431
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   432
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   433
lemma quicksort_outer_remains:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   434
  assumes "crel (quicksort a l r) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   435
   shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   436
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   437
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
   438
  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
   439
  note cr = `crel (quicksort a l r) h h' rs`
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   440
  thus ?case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   441
  proof (cases "r > l")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   442
    case False
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   443
    with cr have "h' = h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   444
      unfolding quicksort.simps [of a l r]
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   445
      by (elim crel_if crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   446
    thus ?thesis by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   447
  next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   448
  case True
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   449
   { 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   450
      fix h1 h2 p ret1 ret2 i
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   451
      assume part: "crel (partition a l r) h h1 p"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   452
      assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ret1"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   453
      assume qs2: "crel (quicksort a (p + 1) r) h2 h' ret2"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   454
      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
   455
      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
   456
      from  partition_outer_remains [OF part True] i_outer
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   457
      have "get_array a h !i = get_array a h1 ! i" by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   458
      moreover
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   459
      with 1(1) [OF True pivot qs1] pivot i_outer
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   460
      have "get_array a h1 ! i = get_array a h2 ! i" by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   461
      moreover
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   462
      with qs2 1(2) [of p h2 h' ret2] True pivot i_outer
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   463
      have "get_array a h2 ! i = get_array a h' ! i" by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   464
      ultimately have "get_array a h ! i= get_array a h' ! i" by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   465
    }
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   466
    with cr show ?thesis
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   467
      unfolding quicksort.simps [of a l r]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   468
      by (elim crel_if crelE crel_assert crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   469
  qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   470
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   471
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   472
lemma quicksort_is_skip:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   473
  assumes "crel (quicksort a l r) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   474
  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
   475
  using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   476
  unfolding quicksort.simps [of a l r]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   477
  by (elim crel_if crel_return) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   478
 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   479
lemma quicksort_sorts:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   480
  assumes "crel (quicksort a l r) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   481
  assumes l_r_length: "l < Heap.length a h" "r < Heap.length a h" 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   482
  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
   483
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   484
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
   485
  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
   486
  note cr = `crel (quicksort a l r) h h' rs`
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   487
  thus ?case
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   488
  proof (cases "r > l")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   489
    case False
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   490
    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
   491
    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
   492
      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
   493
  next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   494
    case True
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   495
    { 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   496
      fix h1 h2 p
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   497
      assume part: "crel (partition a l r) h h1 p"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   498
      assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ()"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   499
      assume qs2: "crel (quicksort a (p + 1) r) h2 h' ()"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   500
      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
   501
      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
   502
     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
   503
      from quicksort_outer_remains [OF qs2] quicksort_outer_remains [OF qs1] pivot quicksort_is_skip[OF qs1]
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   504
      have pivot_unchanged: "get_array a h1 ! p = get_array a h' ! p" by (cases p, auto)
28013
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   505
        (*-- 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
   506
      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
   507
      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
   508
      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
   509
      have left_subarray_remains: "subarray l p a h2 = subarray l p a h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   510
	by (simp add: subarray_eq_samelength_iff)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   511
      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
   512
      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
   513
      have IH2: "sorted (subarray (p + 1) (r + 1) a h')"
28013
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   514
        by (cases "Suc p \<le> r", auto simp add: subarray_Nil)
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   515
           (* -- Secondly, both sublists remain partitioned. *)
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   516
      from partition_partitions[OF part True]
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   517
      have part_conds1: "\<forall>j. j \<in> set (subarray l p a h1) \<longrightarrow> j \<le> get_array a h1 ! p "
28013
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   518
        and part_conds2: "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h1) \<longrightarrow> get_array a h1 ! p \<le> j"
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   519
        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
   520
      from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True
28013
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   521
        length_remains 1(5) pivot multiset_of_sublist [of l p "get_array a h1" "get_array a h2"]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   522
      have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   523
	unfolding Heap.length_def subarray_def by (cases p, auto)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   524
      with left_subarray_remains part_conds1 pivot_unchanged
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   525
      have part_conds2': "\<forall>j. j \<in> set (subarray l p a h') \<longrightarrow> j \<le> get_array a h' ! p"
28013
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   526
        by (simp, subst set_of_multiset_of[symmetric], simp)
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   527
          (* -- 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
   528
      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
   529
      have right_subarray_remains: "subarray (p + 1) (r + 1) a h1 = subarray (p + 1) (r + 1) a h2"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   530
	by (auto simp add: subarray_eq_samelength_iff)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   531
      from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True
28013
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   532
        length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "get_array a h2" "get_array a h'"]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   533
      have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)"
28013
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   534
        unfolding Heap.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
   535
      with right_subarray_remains part_conds2 pivot_unchanged
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   536
      have part_conds1': "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h') \<longrightarrow> get_array a h' ! p \<le> j"
28013
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   537
        by (simp, subst set_of_multiset_of[symmetric], simp)
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   538
          (* -- Thirdly and finally, we show that the array is sorted
e892cedcd638 untabification
haftmann
parents: 27707
diff changeset
   539
          following from the facts above. *)
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   540
      from True pivot 1(5) length_remains have "subarray l (r + 1) a h' = subarray l p a h' @ [get_array a h' ! p] @ subarray (p + 1) (r + 1) a h'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   541
	by (simp add: subarray_nth_array_Cons, cases "l < p") (auto simp add: subarray_append subarray_Nil)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   542
      with IH1' IH2 part_conds1' part_conds2' pivot have ?thesis
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   543
	unfolding subarray_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   544
	apply (auto simp add: sorted_append sorted_Cons all_in_set_sublist'_conv)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   545
	by (auto simp add: set_sublist' dest: le_trans [of _ "get_array a h' ! p"])
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   546
    }
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   547
    with True cr show ?thesis
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   548
      unfolding quicksort.simps [of a l r]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   549
      by (elim crel_if crel_return crelE crel_assert) auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   550
  qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   551
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   552
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   553
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   554
lemma quicksort_is_sort:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   555
  assumes crel: "crel (quicksort a 0 (Heap.length a h - 1)) h h' rs"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   556
  shows "get_array a h' = sort (get_array a h)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   557
proof (cases "get_array a h = []")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   558
  case True
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   559
  with quicksort_is_skip[OF crel] show ?thesis
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   560
  unfolding Heap.length_def by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   561
next
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   562
  case False
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   563
  from quicksort_sorts [OF crel] False have "sorted (sublist' 0 (List.length (get_array a h)) (get_array a h'))"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   564
    unfolding Heap.length_def subarray_def by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   565
  with length_remains[OF crel] have "sorted (get_array a h')"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   566
    unfolding Heap.length_def by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   567
  with quicksort_permutes [OF crel] properties_for_sort show ?thesis by fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   568
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   569
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   570
subsection {* No Errors in quicksort *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   571
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
   572
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
   573
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   574
lemma noError_part1: 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   575
  assumes "l < Heap.length a h" "r < Heap.length a h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   576
  shows "noError (part1 a l r p) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   577
  using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   578
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
   579
  case (1 a l r p)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   580
  thus ?case
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   581
    unfolding part1.simps [of a l r] swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   582
    by (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd elim!: crelE crel_upd crel_nth crel_return)
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
lemma noError_partition:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   586
  assumes "l < r" "l < Heap.length a h" "r < Heap.length a h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   587
  shows "noError (partition a l r) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   588
using assms
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   589
unfolding partition.simps swap_def
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   590
apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_part1 elim!: crelE crel_upd crel_nth crel_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   591
apply (frule part_length_remains)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   592
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
   593
apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   594
apply (frule part_length_remains)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   595
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
   596
apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   597
apply (frule part_length_remains)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   598
apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   599
done
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   600
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   601
lemma noError_quicksort:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   602
  assumes "l < Heap.length a h" "r < Heap.length a h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   603
  shows "noError (quicksort a l r) h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   604
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   605
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
   606
  case (1 a l ri h)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   607
  thus ?case
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   608
    unfolding quicksort.simps [of a l ri]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   609
    apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_assert noError_partition)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   610
    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
   611
    apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   612
    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
   613
    apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   614
    apply (auto elim!: crel_assert dest!: partition_length_remains length_remains)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   615
    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
   616
    apply (erule disjE)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   617
    apply auto
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 28013
diff changeset
   618
    unfolding quicksort.simps [of a "Suc ri" ri]
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   619
    apply (auto intro!: noError_if noError_return)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   620
    done
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   621
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   622
27674
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   623
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   624
subsection {* Example *}
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   625
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   626
definition "qsort a = do
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   627
    k \<leftarrow> length a;
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   628
    quicksort a 0 (k - 1);
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   629
    return a
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   630
  done"
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   631
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   632
ML {* @{code qsort} (Array.fromList [42, 2, 3, 5, 0, 1705, 8, 3, 15]) () *}
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   633
29793
86cac1fab613 changed name space policy for Haskell includes
haftmann
parents: 29399
diff changeset
   634
export_code qsort in SML_imp module_name QSort
86cac1fab613 changed name space policy for Haskell includes
haftmann
parents: 29399
diff changeset
   635
export_code qsort in OCaml module_name QSort file -
86cac1fab613 changed name space policy for Haskell includes
haftmann
parents: 29399
diff changeset
   636
export_code qsort in OCaml_imp module_name QSort file -
86cac1fab613 changed name space policy for Haskell includes
haftmann
parents: 29399
diff changeset
   637
export_code qsort in Haskell module_name QSort file -
27674
2736967f27fd added code generation
haftmann
parents: 27656
diff changeset
   638
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
   639
end