| 69252 |      1 | (*  Title:      HOL/ex/Sorting_Algorithms_Examples.thy
 | 
|  |      2 |     Author:     Florian Haftmann, TU Muenchen
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
|  |      5 | theory Sorting_Algorithms_Examples
 | 
|  |      6 |   imports Main "HOL-Library.Sorting_Algorithms"
 | 
|  |      7 | begin
 | 
|  |      8 | 
 | 
|  |      9 | subsection \<open>Evaluation examples\<close>
 | 
|  |     10 | 
 | 
|  |     11 | definition int_abs_reversed :: "int comparator"
 | 
|  |     12 |   where "int_abs_reversed = key abs (reversed default)"
 | 
|  |     13 | 
 | 
|  |     14 | definition example_1 :: "int list"
 | 
|  |     15 |   where "example_1 = [65, 1705, -2322, 734, 4, (-17::int)]"
 | 
|  |     16 | 
 | 
|  |     17 | definition example_2 :: "int list"
 | 
|  |     18 |   where "example_2 = [-3000..3000]"
 | 
|  |     19 | 
 | 
|  |     20 | ML \<open>
 | 
|  |     21 | local
 | 
|  |     22 | 
 | 
| 69597 |     23 |   val term_of_int_list = HOLogic.mk_list \<^typ>\<open>int\<close>
 | 
|  |     24 |     o map (HOLogic.mk_number \<^typ>\<open>int\<close> o @{code integer_of_int});
 | 
| 69252 |     25 | 
 | 
| 74582 |     26 |   fun raw_sort (ctxt, ct, ks) =
 | 
|  |     27 |     \<^instantiate>\<open>x = ct and y = \<open>Thm.cterm_of ctxt (term_of_int_list ks)\<close>
 | 
| 74584 |     28 |       in cterm \<open>x \<equiv> y\<close> for x y :: \<open>int list\<close>\<close>;
 | 
| 69252 |     29 | 
 | 
|  |     30 |   val (_, sort_oracle) = Context.>>> (Context.map_theory_result
 | 
| 69597 |     31 |     (Thm.add_oracle (\<^binding>\<open>sort\<close>, raw_sort)));
 | 
| 69252 |     32 | 
 | 
|  |     33 | in
 | 
|  |     34 | 
 | 
|  |     35 |   val sort_int_abs_reversed_conv = @{computation_conv "int list" terms: int_abs_reversed
 | 
|  |     36 |     "sort :: int comparator \<Rightarrow> _"
 | 
|  |     37 |     "quicksort :: int comparator \<Rightarrow> _"
 | 
|  |     38 |     "mergesort :: int comparator \<Rightarrow> _"
 | 
|  |     39 |     example_1 example_2
 | 
|  |     40 |   } (fn ctxt => fn ct => fn ks => sort_oracle (ctxt, ks, ct))
 | 
|  |     41 | 
 | 
|  |     42 | end
 | 
|  |     43 | \<close>
 | 
|  |     44 | 
 | 
|  |     45 | declare [[code_timing]]
 | 
|  |     46 | 
 | 
| 69597 |     47 | ML_val \<open>sort_int_abs_reversed_conv \<^context>
 | 
|  |     48 |   \<^cterm>\<open>sort int_abs_reversed example_1\<close>\<close>
 | 
| 69252 |     49 | 
 | 
| 69597 |     50 | ML_val \<open>sort_int_abs_reversed_conv \<^context>
 | 
|  |     51 |   \<^cterm>\<open>quicksort int_abs_reversed example_1\<close>\<close>
 | 
| 69252 |     52 | 
 | 
| 69597 |     53 | ML_val \<open>sort_int_abs_reversed_conv \<^context>
 | 
|  |     54 |   \<^cterm>\<open>mergesort int_abs_reversed example_1\<close>\<close>
 | 
| 69252 |     55 | 
 | 
| 69597 |     56 | ML_val \<open>sort_int_abs_reversed_conv \<^context>
 | 
|  |     57 |   \<^cterm>\<open>sort int_abs_reversed example_2\<close>\<close>
 | 
| 69252 |     58 | 
 | 
| 69597 |     59 | ML_val \<open>sort_int_abs_reversed_conv \<^context>
 | 
|  |     60 |   \<^cterm>\<open>quicksort int_abs_reversed example_2\<close>\<close>
 | 
| 69252 |     61 | 
 | 
| 69597 |     62 | ML_val \<open>sort_int_abs_reversed_conv \<^context>
 | 
|  |     63 |   \<^cterm>\<open>mergesort int_abs_reversed example_2\<close>\<close>
 | 
| 69252 |     64 | 
 | 
|  |     65 | end
 |