| author | wenzelm | 
| Sun, 21 Feb 2010 22:35:02 +0100 | |
| changeset 35262 | 9ea4445d2ccf | 
| parent 35092 | cfe605c54e50 | 
| child 35402 | 115a5a95710a | 
| permissions | -rw-r--r-- | 
| 31775 | 1 | (* Title: HOL/Tools/Function/lexicographic_order.ML | 
| 21131 | 2 | Author: Lukas Bulwahn, TU Muenchen | 
| 3 | ||
| 4 | Method for termination proofs with lexicographic orderings. | |
| 5 | *) | |
| 6 | ||
| 7 | signature LEXICOGRAPHIC_ORDER = | |
| 8 | sig | |
| 33351 | 9 | val lex_order_tac : bool -> Proof.context -> tactic -> tactic | 
| 10 | val lexicographic_order_tac : bool -> Proof.context -> tactic | |
| 30493 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 wenzelm parents: 
30304diff
changeset | 11 | val lexicographic_order : Proof.context -> Proof.method | 
| 21510 
7e72185e4b24
exported mk_base_funs for use by size-change tools
 krauss parents: 
21319diff
changeset | 12 | |
| 21237 | 13 | val setup: theory -> theory | 
| 21131 | 14 | end | 
| 15 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
32952diff
changeset | 16 | structure Lexicographic_Order : LEXICOGRAPHIC_ORDER = | 
| 21131 | 17 | struct | 
| 18 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
32952diff
changeset | 19 | open Function_Lib | 
| 27790 
37b4e65d1dbf
FundefLib.try_proof : attempt a proof and see if it works
 krauss parents: 
27721diff
changeset | 20 | |
| 23074 | 21 | (** General stuff **) | 
| 22 | ||
| 23 | fun mk_measures domT mfuns = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 24 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 25 | val relT = HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 26 | val mlexT = (domT --> HOLogic.natT) --> relT --> relT | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 27 |     fun mk_ms [] = Const (@{const_name Set.empty}, relT)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 28 | | mk_ms (f::fs) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 29 |         Const (@{const_name mlex_prod}, mlexT) $ f $ mk_ms fs
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 30 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 31 | mk_ms mfuns | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 32 | end | 
| 23074 | 33 | |
| 34 | fun del_index n [] = [] | |
| 35 | | del_index n (x :: xs) = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 36 | if n > 0 then x :: del_index (n - 1) xs else xs | 
| 21131 | 37 | |
| 38 | fun transpose ([]::_) = [] | |
| 39 | | transpose xss = map hd xss :: transpose (map tl xss) | |
| 40 | ||
| 23074 | 41 | (** Matrix cell datatype **) | 
| 42 | ||
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 43 | datatype cell = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 44 | Less of thm| LessEq of (thm * thm) | None of (thm * thm) | False of thm; | 
| 23633 | 45 | |
| 23074 | 46 | fun is_Less (Less _) = true | 
| 47 | | is_Less _ = false | |
| 23633 | 48 | |
| 23074 | 49 | fun is_LessEq (LessEq _) = true | 
| 50 | | is_LessEq _ = false | |
| 23633 | 51 | |
| 23437 | 52 | fun pr_cell (Less _ ) = " < " | 
| 23633 | 53 | | pr_cell (LessEq _) = " <=" | 
| 23437 | 54 | | pr_cell (None _) = " ? " | 
| 55 | | pr_cell (False _) = " F " | |
| 23074 | 56 | |
| 57 | ||
| 58 | (** Proof attempts to build the matrix **) | |
| 23633 | 59 | |
| 21131 | 60 | fun dest_term (t : term) = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 61 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 62 | val (vars, prop) = Function_Lib.dest_all_all t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 63 | val (prems, concl) = Logic.strip_horn prop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 64 | val (lhs, rhs) = concl | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 65 | |> HOLogic.dest_Trueprop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 66 | |> HOLogic.dest_mem |> fst | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 67 | |> HOLogic.dest_prod | 
| 
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 | (vars, prems, lhs, rhs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 70 | end | 
| 23633 | 71 | |
| 21131 | 72 | fun mk_goal (vars, prems, lhs, rhs) rel = | 
| 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 concl = HOLogic.mk_binrel rel (lhs, rhs) |> HOLogic.mk_Trueprop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 75 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 76 | fold_rev Logic.all vars (Logic.list_implies (prems, concl)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 77 | end | 
| 23633 | 78 | |
| 79 | fun mk_cell (thy : theory) solve_tac (vars, prems, lhs, rhs) mfun = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 80 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 81 | 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 | 82 | in | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
34974diff
changeset | 83 |     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 | 84 | Solved thm => Less thm | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 85 | | Stuck thm => | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
34974diff
changeset | 86 |       (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 | 87 | Solved thm2 => LessEq (thm2, thm) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 88 | | Stuck thm2 => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 89 | 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 | 90 | else None (thm2, thm) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 91 | | _ => raise Match) (* FIXME *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 92 | | _ => raise Match | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 93 | end | 
| 23074 | 94 | |
| 95 | ||
| 96 | (** Search algorithms **) | |
| 22309 | 97 | |
| 23074 | 98 | fun check_col ls = forall (fn c => is_Less c orelse is_LessEq c) ls andalso not (forall (is_LessEq) ls) | 
| 22309 | 99 | |
| 23074 | 100 | fun transform_table table col = table |> filter_out (fn x => is_Less (nth x col)) |> map (del_index col) | 
| 22309 | 101 | |
| 23074 | 102 | fun transform_order col order = map (fn x => if x >= col then x + 1 else x) order | 
| 23633 | 103 | |
| 21131 | 104 | (* simple depth-first search algorithm for the table *) | 
| 105 | fun search_table table = | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 106 | case table of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 107 | [] => SOME [] | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 108 | | _ => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 109 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 110 | val col = find_index (check_col) (transpose table) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 111 | in case col of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 112 | ~1 => NONE | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 113 | | _ => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 114 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 115 | val order_opt = (table, col) |-> transform_table |> search_table | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 116 | in case order_opt of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 117 | NONE => NONE | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 118 | | SOME order =>SOME (col :: transform_order col order) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 119 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 120 | end | 
| 22258 
0967b03844b5
changes in lexicographic_order termination tactic
 bulwahn parents: 
21817diff
changeset | 121 | |
| 23074 | 122 | (** Proof Reconstruction **) | 
| 123 | ||
| 124 | (* prove row :: cell list -> tactic *) | |
| 125 | fun prove_row (Less less_thm :: _) = | |
| 24576 
32ddd902b0ad
lexicographic_order method uses "<*mlex*>" instead of "measures" => no longer depends on List.thy
 krauss parents: 
23881diff
changeset | 126 |     (rtac @{thm "mlex_less"} 1)
 | 
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24961diff
changeset | 127 | THEN PRIMITIVE (Thm.elim_implies less_thm) | 
| 23437 | 128 | | prove_row (LessEq (lesseq_thm, _) :: tail) = | 
| 24576 
32ddd902b0ad
lexicographic_order method uses "<*mlex*>" instead of "measures" => no longer depends on List.thy
 krauss parents: 
23881diff
changeset | 129 |     (rtac @{thm "mlex_leq"} 1)
 | 
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24961diff
changeset | 130 | THEN PRIMITIVE (Thm.elim_implies lesseq_thm) | 
| 23074 | 131 | THEN prove_row tail | 
| 132 | | prove_row _ = sys_error "lexicographic_order" | |
| 133 | ||
| 134 | ||
| 135 | (** Error reporting **) | |
| 136 | ||
| 23633 | 137 | fun pr_goals ctxt st = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 138 |   Goal_Display.pretty_goals ctxt {total = true, main = false, maxgoals = Thm.nprems_of st} st
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 139 | |> Pretty.chunks | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 140 | |> Pretty.string_of | 
| 23437 | 141 | |
| 142 | fun row_index i = chr (i + 97) | |
| 143 | fun col_index j = string_of_int (j + 1) | |
| 144 | ||
| 23633 | 145 | fun pr_unprovable_cell _ ((i,j), Less _) = "" | 
| 146 | | pr_unprovable_cell ctxt ((i,j), LessEq (_, st)) = | |
| 147 |       "(" ^ 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 | 148 | | pr_unprovable_cell ctxt ((i,j), None (st_leq, st_less)) = | 
| 23633 | 149 |       "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st_less
 | 
| 150 |       ^ "\n(" ^ row_index i ^ ", " ^ col_index j ^ ", <=):\n" ^ pr_goals ctxt st_leq
 | |
| 151 | | pr_unprovable_cell ctxt ((i,j), False st) = | |
| 152 |       "(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^ pr_goals ctxt st
 | |
| 23437 | 153 | |
| 23633 | 154 | fun pr_unprovable_subgoals ctxt table = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 155 | table | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 156 | |> 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 | 157 | |> flat | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 158 | |> map (pr_unprovable_cell ctxt) | 
| 23437 | 159 | |
| 23633 | 160 | fun no_order_msg ctxt table tl measure_funs = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 161 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 162 | val prterm = Syntax.string_of_term ctxt | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 163 | fun pr_fun t i = string_of_int i ^ ") " ^ prterm t | 
| 23074 | 164 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 165 | fun pr_goal t i = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 166 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 167 | val (_, _, lhs, rhs) = dest_term t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 168 | in (* also show prems? *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 169 | i ^ ") " ^ prterm rhs ^ " ~> " ^ prterm lhs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 170 | end | 
| 23074 | 171 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 172 | 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 | 173 | val mc = 1 upto length measure_funs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 174 |     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 | 175 | :: 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 | 176 | val gstr = "Calls:" :: map2 (prefix " " oo pr_goal) tl gc | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 177 | 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 | 178 | val ustr = "Unfinished subgoals:" :: pr_unprovable_subgoals ctxt table | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 179 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 180 | cat_lines (ustr @ gstr @ mstr @ tstr @ | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 181 | ["", "Could not find lexicographic termination order."]) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 182 | end | 
| 23633 | 183 | |
| 23074 | 184 | (** The Main Function **) | 
| 30493 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 wenzelm parents: 
30304diff
changeset | 185 | |
| 33351 | 186 | 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 | 187 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 188 | val thy = ProofContext.theory_of ctxt | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 189 | val ((_ $ (_ $ rel)) :: tl) = prems_of st | 
| 23074 | 190 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 191 | val (domT, _) = HOLogic.dest_prodT (HOLogic.dest_setT (fastype_of rel)) | 
| 23074 | 192 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 193 | val measure_funs = (* 1: generate measures *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 194 | MeasureFunctions.get_measure_functions ctxt domT | 
| 23633 | 195 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 196 | val table = (* 2: create table *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 197 | Par_List.map (fn t => Par_List.map (mk_cell thy solve_tac (dest_term t)) measure_funs) tl | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 198 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 199 | case search_table table of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 200 | 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 | 201 | | SOME order => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 202 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 203 | 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 | 204 | val relation = mk_measures domT (map (nth measure_funs) order) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 205 | val _ = if not quiet | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 206 |           then writeln ("Found termination order: " ^ quote (Syntax.string_of_term ctxt relation))
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 207 | else () | 
| 23074 | 208 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 209 | in (* 4: proof reconstruction *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 210 | 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 | 211 |         THEN (REPEAT (rtac @{thm "wf_mlex"} 1))
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 212 |         THEN (rtac @{thm "wf_empty"} 1)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 213 | THEN EVERY (map prove_row clean_table)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 214 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 215 | end | 
| 21131 | 216 | |
| 33351 | 217 | fun lexicographic_order_tac quiet ctxt = | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
32952diff
changeset | 218 | TRY (Function_Common.apply_termination_rule ctxt 1) | 
| 33351 | 219 | THEN lex_order_tac quiet ctxt | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
32952diff
changeset | 220 | (auto_tac (clasimpset_of ctxt addsimps2 Function_Common.Termination_Simps.get ctxt)) | 
| 30493 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 wenzelm parents: 
30304diff
changeset | 221 | |
| 33351 | 222 | val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac false | 
| 21201 | 223 | |
| 30541 | 224 | val setup = | 
| 225 |   Method.setup @{binding lexicographic_order}
 | |
| 226 | (Method.sections clasimp_modifiers >> (K lexicographic_order)) | |
| 227 | "termination prover for lexicographic orderings" | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
32952diff
changeset | 228 | #> Context.theory_map (Function_Common.set_termination_prover lexicographic_order) | 
| 21131 | 229 | |
| 30493 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 wenzelm parents: 
30304diff
changeset | 230 | end; |