| author | wenzelm | 
| Mon, 13 Jan 2014 20:20:44 +0100 | |
| changeset 55007 | 0c07990363a3 | 
| parent 51958 | bca32217b304 | 
| child 56231 | b98813774a63 | 
| 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  | 
|
| 21237 | 12  | 
val setup: theory -> theory  | 
| 21131 | 13  | 
end  | 
14  | 
||
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
32952 
diff
changeset
 | 
15  | 
structure Lexicographic_Order : LEXICOGRAPHIC_ORDER =  | 
| 21131 | 16  | 
struct  | 
17  | 
||
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
32952 
diff
changeset
 | 
18  | 
open Function_Lib  | 
| 
27790
 
37b4e65d1dbf
FundefLib.try_proof : attempt a proof and see if it works
 
krauss 
parents: 
27721 
diff
changeset
 | 
19  | 
|
| 23074 | 20  | 
(** General stuff **)  | 
21  | 
||
22  | 
fun mk_measures domT mfuns =  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
23  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
24  | 
val relT = HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT))  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
25  | 
val mlexT = (domT --> HOLogic.natT) --> relT --> relT  | 
| 35402 | 26  | 
    fun mk_ms [] = Const (@{const_abbrev Set.empty}, relT)
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
27  | 
| mk_ms (f::fs) =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
28  | 
        Const (@{const_name mlex_prod}, mlexT) $ f $ mk_ms fs
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
29  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
30  | 
mk_ms mfuns  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
31  | 
end  | 
| 23074 | 32  | 
|
33  | 
fun del_index n [] = []  | 
|
34  | 
| del_index n (x :: xs) =  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
35  | 
if n > 0 then x :: del_index (n - 1) xs else xs  | 
| 21131 | 36  | 
|
37  | 
fun transpose ([]::_) = []  | 
|
38  | 
| transpose xss = map hd xss :: transpose (map tl xss)  | 
|
39  | 
||
| 23074 | 40  | 
(** Matrix cell datatype **)  | 
41  | 
||
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
42  | 
datatype cell =  | 
| 
39926
 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 
krauss 
parents: 
39125 
diff
changeset
 | 
43  | 
Less of thm | LessEq of (thm * thm) | None of (thm * thm) | False of thm;  | 
| 23633 | 44  | 
|
| 
39926
 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 
krauss 
parents: 
39125 
diff
changeset
 | 
45  | 
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: 
39125 
diff
changeset
 | 
46  | 
fun is_LessEq lcell = case Lazy.force lcell of LessEq _ => true | _ => false;  | 
| 23633 | 47  | 
|
| 23074 | 48  | 
|
49  | 
(** Proof attempts to build the matrix **)  | 
|
| 23633 | 50  | 
|
| 39928 | 51  | 
fun dest_term t =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
52  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
53  | 
val (vars, prop) = Function_Lib.dest_all_all t  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
54  | 
val (prems, concl) = Logic.strip_horn prop  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
55  | 
val (lhs, rhs) = concl  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
56  | 
|> HOLogic.dest_Trueprop  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
57  | 
|> HOLogic.dest_mem |> fst  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
58  | 
|> HOLogic.dest_prod  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
59  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
60  | 
(vars, prems, lhs, rhs)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
61  | 
end  | 
| 23633 | 62  | 
|
| 21131 | 63  | 
fun mk_goal (vars, prems, lhs, rhs) rel =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
64  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
65  | 
val concl = HOLogic.mk_binrel rel (lhs, rhs) |> HOLogic.mk_Trueprop  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
66  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
67  | 
fold_rev Logic.all vars (Logic.list_implies (prems, concl))  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
68  | 
end  | 
| 23633 | 69  | 
|
| 39928 | 70  | 
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: 
33855 
diff
changeset
 | 
71  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
72  | 
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: 
33855 
diff
changeset
 | 
73  | 
in  | 
| 
35092
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
34974 
diff
changeset
 | 
74  | 
    case try_proof (goals @{const_name Orderings.less}) solve_tac of
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
75  | 
Solved thm => Less thm  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
76  | 
| Stuck thm =>  | 
| 
35092
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
34974 
diff
changeset
 | 
77  | 
      (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: 
33855 
diff
changeset
 | 
78  | 
Solved thm2 => LessEq (thm2, thm)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
79  | 
| Stuck thm2 =>  | 
| 45740 | 80  | 
         if prems_of thm2 = [HOLogic.Trueprop $ @{term False}] then False thm2
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
81  | 
else None (thm2, thm)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
82  | 
| _ => raise Match) (* FIXME *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
83  | 
| _ => raise Match  | 
| 
39926
 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 
krauss 
parents: 
39125 
diff
changeset
 | 
84  | 
end);  | 
| 23074 | 85  | 
|
86  | 
||
87  | 
(** Search algorithms **)  | 
|
| 22309 | 88  | 
|
| 39928 | 89  | 
fun check_col ls = forall (fn c => is_Less c orelse is_LessEq c) ls andalso not (forall is_LessEq ls)  | 
| 22309 | 90  | 
|
| 23074 | 91  | 
fun transform_table table col = table |> filter_out (fn x => is_Less (nth x col)) |> map (del_index col)  | 
| 22309 | 92  | 
|
| 23074 | 93  | 
fun transform_order col order = map (fn x => if x >= col then x + 1 else x) order  | 
| 23633 | 94  | 
|
| 21131 | 95  | 
(* simple depth-first search algorithm for the table *)  | 
| 39928 | 96  | 
fun search_table [] = SOME []  | 
97  | 
| search_table table =  | 
|
98  | 
case find_index check_col (transpose table) of  | 
|
99  | 
~1 => NONE  | 
|
100  | 
| col =>  | 
|
101  | 
(case (table, col) |-> transform_table |> search_table of  | 
|
102  | 
NONE => NONE  | 
|
103  | 
| SOME order => SOME (col :: transform_order col order))  | 
|
104  | 
||
| 
22258
 
0967b03844b5
changes in lexicographic_order termination tactic
 
bulwahn 
parents: 
21817 
diff
changeset
 | 
105  | 
|
| 23074 | 106  | 
(** Proof Reconstruction **)  | 
107  | 
||
| 39927 | 108  | 
fun prove_row (c :: cs) =  | 
109  | 
(case Lazy.force c of  | 
|
110  | 
Less thm =>  | 
|
111  | 
       rtac @{thm "mlex_less"} 1
 | 
|
112  | 
THEN PRIMITIVE (Thm.elim_implies thm)  | 
|
113  | 
| LessEq (thm, _) =>  | 
|
114  | 
       rtac @{thm "mlex_leq"} 1
 | 
|
115  | 
THEN PRIMITIVE (Thm.elim_implies thm)  | 
|
116  | 
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: 
39928 
diff
changeset
 | 
117  | 
| _ => raise General.Fail "lexicographic_order")  | 
| 39927 | 118  | 
| prove_row [] = no_tac;  | 
| 23074 | 119  | 
|
120  | 
||
121  | 
(** Error reporting **)  | 
|
122  | 
||
| 49856 | 123  | 
fun row_index i = chr (i + 97) (* FIXME not scalable wrt. chr range *)  | 
124  | 
fun col_index j = string_of_int (j + 1) (* FIXME not scalable wrt. table layout *)  | 
|
| 23437 | 125  | 
|
| 49856 | 126  | 
fun pr_unprovable_cell _ ((i,j), Less _) = []  | 
| 23633 | 127  | 
| pr_unprovable_cell ctxt ((i,j), LessEq (_, st)) =  | 
| 
51958
 
bca32217b304
retain goal display options when printing error messages, to avoid breakdown for huge goals;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
128  | 
      ["(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^
 | 
| 
 
bca32217b304
retain goal display options when printing error messages, to avoid breakdown for huge goals;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
129  | 
Goal_Display.string_of_goal ctxt st]  | 
| 
26875
 
e18574413bc4
Measure functions can now be declared via special rules, allowing for a
 
krauss 
parents: 
26749 
diff
changeset
 | 
130  | 
| pr_unprovable_cell ctxt ((i,j), None (st_leq, st_less)) =  | 
| 
51958
 
bca32217b304
retain goal display options when printing error messages, to avoid breakdown for huge goals;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
131  | 
      ["(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^
 | 
| 
 
bca32217b304
retain goal display options when printing error messages, to avoid breakdown for huge goals;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
132  | 
Goal_Display.string_of_goal ctxt st_less ^  | 
| 
 
bca32217b304
retain goal display options when printing error messages, to avoid breakdown for huge goals;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
133  | 
       "\n(" ^ row_index i ^ ", " ^ col_index j ^ ", <=):\n" ^
 | 
| 
 
bca32217b304
retain goal display options when printing error messages, to avoid breakdown for huge goals;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
134  | 
Goal_Display.string_of_goal ctxt st_leq]  | 
| 23633 | 135  | 
| pr_unprovable_cell ctxt ((i,j), False st) =  | 
| 
51958
 
bca32217b304
retain goal display options when printing error messages, to avoid breakdown for huge goals;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
136  | 
      ["(" ^ row_index i ^ ", " ^ col_index j ^ ", <):\n" ^
 | 
| 
 
bca32217b304
retain goal display options when printing error messages, to avoid breakdown for huge goals;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
137  | 
Goal_Display.string_of_goal ctxt st]  | 
| 23437 | 138  | 
|
| 23633 | 139  | 
fun pr_unprovable_subgoals ctxt table =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
140  | 
table  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
141  | 
|> 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: 
33855 
diff
changeset
 | 
142  | 
|> flat  | 
| 49856 | 143  | 
|> maps (pr_unprovable_cell ctxt)  | 
| 23437 | 144  | 
|
| 39928 | 145  | 
fun pr_cell (Less _ ) = " < "  | 
146  | 
| pr_cell (LessEq _) = " <="  | 
|
147  | 
| pr_cell (None _) = " ? "  | 
|
148  | 
| pr_cell (False _) = " F "  | 
|
149  | 
||
| 
39926
 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 
krauss 
parents: 
39125 
diff
changeset
 | 
150  | 
fun no_order_msg ctxt ltable tl measure_funs =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
151  | 
let  | 
| 
39926
 
4b3b384d3de3
lexicographic_order: replace parallel map with laziness to avoid wasting cycles on proofs that are not needed
 
krauss 
parents: 
39125 
diff
changeset
 | 
152  | 
val table = map (map Lazy.force) ltable  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
153  | 
val prterm = Syntax.string_of_term ctxt  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
154  | 
fun pr_fun t i = string_of_int i ^ ") " ^ prterm t  | 
| 23074 | 155  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
156  | 
fun pr_goal t i =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
157  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
158  | 
val (_, _, lhs, rhs) = dest_term t  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
159  | 
in (* also show prems? *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
160  | 
i ^ ") " ^ prterm rhs ^ " ~> " ^ prterm lhs  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
161  | 
end  | 
| 23074 | 162  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
163  | 
val gc = map (fn i => chr (i + 96)) (1 upto length table)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
164  | 
val mc = 1 upto length measure_funs  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
165  | 
    val tstr = "Result matrix:" ::  ("   " ^ implode (map (enclose " " " " o string_of_int) mc))
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
166  | 
:: 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: 
33855 
diff
changeset
 | 
167  | 
val gstr = "Calls:" :: map2 (prefix " " oo pr_goal) tl gc  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
168  | 
val mstr = "Measures:" :: map2 (prefix " " oo pr_fun) measure_funs mc  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
169  | 
val ustr = "Unfinished subgoals:" :: pr_unprovable_subgoals ctxt table  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
170  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
171  | 
cat_lines (ustr @ gstr @ mstr @ tstr @  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
172  | 
["", "Could not find lexicographic termination order."])  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
173  | 
end  | 
| 23633 | 174  | 
|
| 23074 | 175  | 
(** The Main Function **)  | 
| 
30493
 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 
wenzelm 
parents: 
30304 
diff
changeset
 | 
176  | 
|
| 33351 | 177  | 
fun lex_order_tac quiet ctxt solve_tac (st: thm) =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
178  | 
let  | 
| 42361 | 179  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
180  | 
val ((_ $ (_ $ rel)) :: tl) = prems_of st  | 
| 23074 | 181  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
182  | 
val (domT, _) = HOLogic.dest_prodT (HOLogic.dest_setT (fastype_of rel))  | 
| 23074 | 183  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
184  | 
val measure_funs = (* 1: generate measures *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
185  | 
MeasureFunctions.get_measure_functions ctxt domT  | 
| 23633 | 186  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
187  | 
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: 
39125 
diff
changeset
 | 
188  | 
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: 
33855 
diff
changeset
 | 
189  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
190  | 
case search_table table of  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
191  | 
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: 
33855 
diff
changeset
 | 
192  | 
| SOME order =>  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
193  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
194  | 
val clean_table = map (fn x => map (nth x) order) table  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
195  | 
val relation = mk_measures domT (map (nth measure_funs) order)  | 
| 
41540
 
414a68d72279
more precise pretty printing -- to accomodate Scala message layout;
 
wenzelm 
parents: 
40317 
diff
changeset
 | 
196  | 
val _ =  | 
| 
 
414a68d72279
more precise pretty printing -- to accomodate Scala message layout;
 
wenzelm 
parents: 
40317 
diff
changeset
 | 
197  | 
if not quiet then  | 
| 
 
414a68d72279
more precise pretty printing -- to accomodate Scala message layout;
 
wenzelm 
parents: 
40317 
diff
changeset
 | 
198  | 
Pretty.writeln (Pretty.block  | 
| 
 
414a68d72279
more precise pretty printing -- to accomodate Scala message layout;
 
wenzelm 
parents: 
40317 
diff
changeset
 | 
199  | 
[Pretty.str "Found termination order:", Pretty.brk 1,  | 
| 
 
414a68d72279
more precise pretty printing -- to accomodate Scala message layout;
 
wenzelm 
parents: 
40317 
diff
changeset
 | 
200  | 
Pretty.quote (Syntax.pretty_term ctxt relation)])  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
201  | 
else ()  | 
| 23074 | 202  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
203  | 
in (* 4: proof reconstruction *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
204  | 
st |> (PRIMITIVE (cterm_instantiate [(cterm_of thy rel, cterm_of thy relation)])  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
205  | 
        THEN (REPEAT (rtac @{thm "wf_mlex"} 1))
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
206  | 
        THEN (rtac @{thm "wf_empty"} 1)
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
207  | 
THEN EVERY (map prove_row clean_table))  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
208  | 
end  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
33855 
diff
changeset
 | 
209  | 
end  | 
| 21131 | 210  | 
|
| 33351 | 211  | 
fun lexicographic_order_tac quiet ctxt =  | 
| 42793 | 212  | 
TRY (Function_Common.apply_termination_rule ctxt 1) THEN  | 
213  | 
lex_order_tac quiet ctxt  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49856 
diff
changeset
 | 
214  | 
(auto_tac (ctxt addsimps Function_Common.Termination_Simps.get ctxt))  | 
| 
30493
 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 
wenzelm 
parents: 
30304 
diff
changeset
 | 
215  | 
|
| 30541 | 216  | 
val setup =  | 
| 47432 | 217  | 
Context.theory_map (Function_Common.set_termination_prover (lexicographic_order_tac false))  | 
| 21131 | 218  | 
|
| 
30493
 
b8570ea1ce25
provide regular ML interfaces for Isar source language elements;
 
wenzelm 
parents: 
30304 
diff
changeset
 | 
219  | 
end;  |