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