| author | huffman | 
| Wed, 09 Nov 2011 15:33:24 +0100 | |
| changeset 45437 | 958d19d3405b | 
| parent 42795 | 66fcc9882784 | 
| child 45740 | 132a3e1c0fe5 | 
| permissions | -rw-r--r-- | 
| 31775 | 1 | (* Title: HOL/Tools/Function/lexicographic_order.ML | 
| 21131 | 2 | Author: Lukas Bulwahn, TU Muenchen | 
| 39928 | 3 | Author: Alexander Krauss, TU Muenchen | 
| 21131 | 4 | |
| 39928 | 5 | Termination proofs with lexicographic orders. | 
| 21131 | 6 | *) | 
| 7 | ||
| 8 | signature LEXICOGRAPHIC_ORDER = | |
| 9 | sig | |
| 33351 | 10 | val lex_order_tac : bool -> Proof.context -> tactic -> tactic | 
| 11 | val lexicographic_order_tac : bool -> Proof.context -> tactic | |
| 30493 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 wenzelm parents: 
30304diff
changeset | 12 | val lexicographic_order : Proof.context -> Proof.method | 
| 21510 
7e72185e4b24
exported mk_base_funs for use by size-change tools
 krauss parents: 
21319diff
changeset | 13 | |
| 21237 | 14 | val setup: theory -> theory | 
| 21131 | 15 | end | 
| 16 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
32952diff
changeset | 17 | structure Lexicographic_Order : LEXICOGRAPHIC_ORDER = | 
| 21131 | 18 | struct | 
| 19 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
32952diff
changeset | 20 | open Function_Lib | 
| 27790 
37b4e65d1dbf
FundefLib.try_proof : attempt a proof and see if it works
 krauss parents: 
27721diff
changeset | 21 | |
| 23074 | 22 | (** General stuff **) | 
| 23 | ||
| 24 | fun mk_measures domT mfuns = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 25 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 26 | val relT = HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 27 | val mlexT = (domT --> HOLogic.natT) --> relT --> relT | 
| 35402 | 28 |     fun mk_ms [] = Const (@{const_abbrev Set.empty}, relT)
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 29 | | mk_ms (f::fs) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 30 |         Const (@{const_name mlex_prod}, mlexT) $ f $ mk_ms fs
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 31 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 32 | mk_ms mfuns | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 33 | end | 
| 23074 | 34 | |
| 35 | fun del_index n [] = [] | |
| 36 | | del_index n (x :: xs) = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 37 | if n > 0 then x :: del_index (n - 1) xs else xs | 
| 21131 | 38 | |
| 39 | fun transpose ([]::_) = [] | |
| 40 | | transpose xss = map hd xss :: transpose (map tl xss) | |
| 41 | ||
| 23074 | 42 | (** Matrix cell datatype **) | 
| 43 | ||
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 44 | datatype cell = | 
| 39926 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 krauss parents: 
39125diff
changeset | 45 | Less of thm | LessEq of (thm * thm) | None of (thm * thm) | False of thm; | 
| 23633 | 46 | |
| 39926 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 krauss parents: 
39125diff
changeset | 47 | fun is_Less lcell = case Lazy.force lcell of Less _ => true | _ => false; | 
| 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 krauss parents: 
39125diff
changeset | 48 | fun is_LessEq lcell = case Lazy.force lcell of LessEq _ => true | _ => false; | 
| 23633 | 49 | |
| 23074 | 50 | |
| 51 | (** Proof attempts to build the matrix **) | |
| 23633 | 52 | |
| 39928 | 53 | fun dest_term t = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 54 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 55 | val (vars, prop) = Function_Lib.dest_all_all t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 56 | val (prems, concl) = Logic.strip_horn prop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 57 | val (lhs, rhs) = concl | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 58 | |> HOLogic.dest_Trueprop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 59 | |> HOLogic.dest_mem |> fst | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 60 | |> HOLogic.dest_prod | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 61 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 62 | (vars, prems, lhs, rhs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 63 | end | 
| 23633 | 64 | |
| 21131 | 65 | fun mk_goal (vars, prems, lhs, rhs) rel = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 66 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 67 | val concl = HOLogic.mk_binrel rel (lhs, rhs) |> HOLogic.mk_Trueprop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 68 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 69 | fold_rev Logic.all vars (Logic.list_implies (prems, concl)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 70 | end | 
| 23633 | 71 | |
| 39928 | 72 | fun mk_cell thy solve_tac (vars, prems, lhs, rhs) mfun = Lazy.lazy (fn _ => | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 73 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 74 | val goals = cterm_of thy o mk_goal (vars, prems, mfun $ lhs, mfun $ rhs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 75 | in | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
34974diff
changeset | 76 |     case try_proof (goals @{const_name Orderings.less}) solve_tac of
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 77 | Solved thm => Less thm | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 78 | | Stuck thm => | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
34974diff
changeset | 79 |       (case try_proof (goals @{const_name Orderings.less_eq}) solve_tac of
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 80 | Solved thm2 => LessEq (thm2, thm) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 81 | | Stuck thm2 => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 82 | if prems_of thm2 = [HOLogic.Trueprop $ HOLogic.false_const] then False thm2 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 83 | else None (thm2, thm) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 84 | | _ => raise Match) (* FIXME *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 85 | | _ => raise Match | 
| 39926 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 krauss parents: 
39125diff
changeset | 86 | end); | 
| 23074 | 87 | |
| 88 | ||
| 89 | (** Search algorithms **) | |
| 22309 | 90 | |
| 39928 | 91 | fun check_col ls = forall (fn c => is_Less c orelse is_LessEq c) ls andalso not (forall is_LessEq ls) | 
| 22309 | 92 | |
| 23074 | 93 | fun transform_table table col = table |> filter_out (fn x => is_Less (nth x col)) |> map (del_index col) | 
| 22309 | 94 | |
| 23074 | 95 | fun transform_order col order = map (fn x => if x >= col then x + 1 else x) order | 
| 23633 | 96 | |
| 21131 | 97 | (* simple depth-first search algorithm for the table *) | 
| 39928 | 98 | fun search_table [] = SOME [] | 
| 99 | | search_table table = | |
| 100 | case find_index check_col (transpose table) of | |
| 101 | ~1 => NONE | |
| 102 | | col => | |
| 103 | (case (table, col) |-> transform_table |> search_table of | |
| 104 | NONE => NONE | |
| 105 | | SOME order => SOME (col :: transform_order col order)) | |
| 106 | ||
| 22258 
0967b03844b5
changes in lexicographic_order termination tactic
 bulwahn parents: 
21817diff
changeset | 107 | |
| 23074 | 108 | (** Proof Reconstruction **) | 
| 109 | ||
| 110 | (* prove row :: cell list -> tactic *) | |
| 39927 | 111 | fun prove_row (c :: cs) = | 
| 112 | (case Lazy.force c of | |
| 113 | Less thm => | |
| 114 |        rtac @{thm "mlex_less"} 1
 | |
| 115 | THEN PRIMITIVE (Thm.elim_implies thm) | |
| 116 | | LessEq (thm, _) => | |
| 117 |        rtac @{thm "mlex_leq"} 1
 | |
| 118 | THEN PRIMITIVE (Thm.elim_implies thm) | |
| 119 | THEN prove_row cs | |
| 40317 
1eac228c52b3
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 wenzelm parents: 
39928diff
changeset | 120 | | _ => raise General.Fail "lexicographic_order") | 
| 39927 | 121 | | prove_row [] = no_tac; | 
| 23074 | 122 | |
| 123 | ||
| 124 | (** Error reporting **) | |
| 125 | ||
| 23633 | 126 | fun pr_goals ctxt st = | 
| 39125 
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
 wenzelm parents: 
36521diff
changeset | 127 | Goal_Display.pretty_goals | 
| 
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
 wenzelm parents: 
36521diff
changeset | 128 | (Config.put Goal_Display.goals_limit (Thm.nprems_of st) ctxt) st | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 129 | |> Pretty.chunks | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 130 | |> Pretty.string_of | 
| 23437 | 131 | |
| 132 | fun row_index i = chr (i + 97) | |
| 133 | fun col_index j = string_of_int (j + 1) | |
| 134 | ||
| 23633 | 135 | fun pr_unprovable_cell _ ((i,j), Less _) = "" | 
| 136 | | pr_unprovable_cell ctxt ((i,j), LessEq (_, st)) = | |
| 137 |       "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st
 | |
| 26875 
e18574413bc4
Measure functions can now be declared via special rules, allowing for a
 krauss parents: 
26749diff
changeset | 138 | | pr_unprovable_cell ctxt ((i,j), None (st_leq, st_less)) = | 
| 23633 | 139 |       "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st_less
 | 
| 140 |       ^ "\n(" ^ row_index i ^ ", " ^ col_index j ^ ", <=):\n" ^ pr_goals ctxt st_leq
 | |
| 141 | | pr_unprovable_cell ctxt ((i,j), False st) = | |
| 142 |       "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st
 | |
| 23437 | 143 | |
| 23633 | 144 | fun pr_unprovable_subgoals ctxt table = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 145 | table | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 146 | |> map_index (fn (i,cs) => map_index (fn (j,x) => ((i,j), x)) cs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 147 | |> flat | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 148 | |> map (pr_unprovable_cell ctxt) | 
| 23437 | 149 | |
| 39928 | 150 | fun pr_cell (Less _ ) = " < " | 
| 151 | | pr_cell (LessEq _) = " <=" | |
| 152 | | pr_cell (None _) = " ? " | |
| 153 | | pr_cell (False _) = " F " | |
| 154 | ||
| 39926 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 krauss parents: 
39125diff
changeset | 155 | fun no_order_msg ctxt ltable tl measure_funs = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 156 | let | 
| 39926 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 krauss parents: 
39125diff
changeset | 157 | val table = map (map Lazy.force) ltable | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 158 | val prterm = Syntax.string_of_term ctxt | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 159 | fun pr_fun t i = string_of_int i ^ ") " ^ prterm t | 
| 23074 | 160 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 161 | fun pr_goal t i = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 162 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 163 | val (_, _, lhs, rhs) = dest_term t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 164 | in (* also show prems? *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 165 | i ^ ") " ^ prterm rhs ^ " ~> " ^ prterm lhs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 166 | end | 
| 23074 | 167 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 168 | val gc = map (fn i => chr (i + 96)) (1 upto length table) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 169 | val mc = 1 upto length measure_funs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 170 |     val tstr = "Result matrix:" ::  ("   " ^ implode (map (enclose " " " " o string_of_int) mc))
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 171 | :: map2 (fn r => fn i => i ^ ": " ^ implode (map pr_cell r)) table gc | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 172 | val gstr = "Calls:" :: map2 (prefix " " oo pr_goal) tl gc | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 173 | val mstr = "Measures:" :: map2 (prefix " " oo pr_fun) measure_funs mc | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 174 | val ustr = "Unfinished subgoals:" :: pr_unprovable_subgoals ctxt table | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 175 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 176 | cat_lines (ustr @ gstr @ mstr @ tstr @ | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 177 | ["", "Could not find lexicographic termination order."]) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 178 | end | 
| 23633 | 179 | |
| 23074 | 180 | (** The Main Function **) | 
| 30493 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 wenzelm parents: 
30304diff
changeset | 181 | |
| 33351 | 182 | fun lex_order_tac quiet ctxt solve_tac (st: thm) = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 183 | let | 
| 42361 | 184 | val thy = Proof_Context.theory_of ctxt | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 185 | val ((_ $ (_ $ rel)) :: tl) = prems_of st | 
| 23074 | 186 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 187 | val (domT, _) = HOLogic.dest_prodT (HOLogic.dest_setT (fastype_of rel)) | 
| 23074 | 188 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 189 | val measure_funs = (* 1: generate measures *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 190 | MeasureFunctions.get_measure_functions ctxt domT | 
| 23633 | 191 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 192 | val table = (* 2: create table *) | 
| 39926 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 krauss parents: 
39125diff
changeset | 193 | map (fn t => map (mk_cell thy solve_tac (dest_term t)) measure_funs) tl | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 194 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 195 | case search_table table of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 196 | NONE => if quiet then no_tac st else error (no_order_msg ctxt table tl measure_funs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 197 | | SOME order => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 198 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 199 | val clean_table = map (fn x => map (nth x) order) table | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 200 | val relation = mk_measures domT (map (nth measure_funs) order) | 
| 41540 
414a68d72279
more precise pretty printing -- to accomodate Scala message layout;
 wenzelm parents: 
40317diff
changeset | 201 | val _ = | 
| 
414a68d72279
more precise pretty printing -- to accomodate Scala message layout;
 wenzelm parents: 
40317diff
changeset | 202 | if not quiet then | 
| 
414a68d72279
more precise pretty printing -- to accomodate Scala message layout;
 wenzelm parents: 
40317diff
changeset | 203 | Pretty.writeln (Pretty.block | 
| 
414a68d72279
more precise pretty printing -- to accomodate Scala message layout;
 wenzelm parents: 
40317diff
changeset | 204 | [Pretty.str "Found termination order:", Pretty.brk 1, | 
| 
414a68d72279
more precise pretty printing -- to accomodate Scala message layout;
 wenzelm parents: 
40317diff
changeset | 205 | Pretty.quote (Syntax.pretty_term ctxt relation)]) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 206 | else () | 
| 23074 | 207 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 208 | in (* 4: proof reconstruction *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 209 | st |> (PRIMITIVE (cterm_instantiate [(cterm_of thy rel, cterm_of thy relation)]) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 210 |         THEN (REPEAT (rtac @{thm "wf_mlex"} 1))
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 211 |         THEN (rtac @{thm "wf_empty"} 1)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 212 | THEN EVERY (map prove_row clean_table)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 213 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 214 | end | 
| 21131 | 215 | |
| 33351 | 216 | fun lexicographic_order_tac quiet ctxt = | 
| 42793 | 217 | TRY (Function_Common.apply_termination_rule ctxt 1) THEN | 
| 218 | lex_order_tac quiet ctxt | |
| 42795 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 219 | (auto_tac (map_simpset (fn ss => ss addsimps Function_Common.Termination_Simps.get ctxt) ctxt)) | 
| 30493 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 wenzelm parents: 
30304diff
changeset | 220 | |
| 33351 | 221 | val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac false | 
| 21201 | 222 | |
| 30541 | 223 | val setup = | 
| 224 |   Method.setup @{binding lexicographic_order}
 | |
| 225 | (Method.sections clasimp_modifiers >> (K lexicographic_order)) | |
| 226 | "termination prover for lexicographic orderings" | |
| 36521 | 227 | #> Context.theory_map (Function_Common.set_termination_prover (lexicographic_order_tac false)) | 
| 21131 | 228 | |
| 30493 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 wenzelm parents: 
30304diff
changeset | 229 | end; |