| author | blanchet | 
| Fri, 23 Aug 2013 16:02:32 +0200 | |
| changeset 53157 | c8369b691d04 | 
| parent 52732 | b4da1f2ec73f | 
| child 54742 | 7a86358a3c0b | 
| permissions | -rw-r--r-- | 
| 31775 | 1 | (* Title: HOL/Tools/Function/termination.ML | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 2 | Author: Alexander Krauss, TU Muenchen | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 3 | |
| 41114 | 4 | Context data for termination proofs. | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 5 | *) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 6 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 7 | signature TERMINATION = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 8 | sig | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 9 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 10 | type data | 
| 39924 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 11 | datatype cell = Less of thm | LessEq of thm * thm | None of thm * thm | False of thm | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 12 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 13 | val mk_sumcases : data -> typ -> term list -> term | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 14 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 15 | val get_num_points : data -> int | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 16 | val get_types : data -> int -> typ | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 17 | val get_measures : data -> int -> term list | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 18 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 19 | val get_chain : data -> term -> term -> thm option option | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 20 | val get_descent : data -> term -> term -> term -> cell option | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 21 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 22 | val dest_call : data -> term -> ((string * typ) list * int * term * int * term * term) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 23 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 24 | val CALLS : (term list * int -> tactic) -> int -> tactic | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 25 | |
| 39924 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 26 | (* Termination tactics *) | 
| 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 27 | type ttac = data -> int -> tactic | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 28 | |
| 39924 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 29 | val TERMINATION : Proof.context -> tactic -> ttac -> int -> tactic | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 30 | |
| 30607 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 wenzelm parents: 
30304diff
changeset | 31 | val wf_union_tac : Proof.context -> tactic | 
| 34228 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 32 | |
| 39924 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 33 | val decompose_tac : ttac | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 34 | end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 35 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 36 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 37 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 38 | structure Termination : TERMINATION = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 39 | struct | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 40 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
32683diff
changeset | 41 | open Function_Lib | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 42 | |
| 35408 | 43 | val term2_ord = prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord | 
| 31971 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 wenzelm parents: 
31775diff
changeset | 44 | structure Term2tab = Table(type key = term * term val ord = term2_ord); | 
| 35408 | 45 | structure Term3tab = | 
| 46 | Table(type key = term * (term * term) val ord = prod_ord Term_Ord.fast_term_ord term2_ord); | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 47 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 48 | (** Analyzing binary trees **) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 49 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 50 | (* Skeleton of a tree structure *) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 51 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 52 | datatype skel = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 53 | SLeaf of int (* index *) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 54 | | SBranch of (skel * skel) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 55 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 56 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 57 | (* abstract make and dest functions *) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 58 | fun mk_tree leaf branch = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 59 | let fun mk (SLeaf i) = leaf i | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 60 | | mk (SBranch (s, t)) = branch (mk s, mk t) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 61 | in mk end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 62 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 63 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 64 | fun dest_tree split = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 65 | let fun dest (SLeaf i) x = [(i, x)] | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 66 | | dest (SBranch (s, t)) x = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 67 | let val (l, r) = split x | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 68 | in dest s l @ dest t r end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 69 | in dest end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 70 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 71 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 72 | (* concrete versions for sum types *) | 
| 32602 | 73 | fun is_inj (Const (@{const_name Sum_Type.Inl}, _) $ _) = true
 | 
| 74 |   | is_inj (Const (@{const_name Sum_Type.Inr}, _) $ _) = true
 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 75 | | is_inj _ = false | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 76 | |
| 32602 | 77 | fun dest_inl (Const (@{const_name Sum_Type.Inl}, _) $ t) = SOME t
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 78 | | dest_inl _ = NONE | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 79 | |
| 32602 | 80 | fun dest_inr (Const (@{const_name Sum_Type.Inr}, _) $ t) = SOME t
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 81 | | dest_inr _ = NONE | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 82 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 83 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 84 | fun mk_skel ps = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 85 | let | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 86 | fun skel i ps = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 87 | if forall is_inj ps andalso not (null ps) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 88 | then let | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 89 | val (j, s) = skel i (map_filter dest_inl ps) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 90 | val (k, t) = skel j (map_filter dest_inr ps) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 91 | in (k, SBranch (s, t)) end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 92 | else (i + 1, SLeaf i) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 93 | in | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 94 | snd (skel 0 ps) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 95 | end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 96 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 97 | (* compute list of types for nodes *) | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37391diff
changeset | 98 | fun node_types sk T = dest_tree (fn Type (@{type_name Sum_Type.sum}, [LT, RT]) => (LT, RT)) sk T |> map snd
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 99 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 100 | (* find index and raw term *) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 101 | fun dest_inj (SLeaf i) trm = (i, trm) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 102 | | dest_inj (SBranch (s, t)) trm = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 103 | case dest_inl trm of | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 104 | SOME trm' => dest_inj s trm' | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 105 | | _ => dest_inj t (the (dest_inr trm)) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 106 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 107 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 108 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 109 | (** Matrix cell datatype **) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 110 | |
| 39924 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 111 | datatype cell = Less of thm | LessEq of thm * thm | None of thm * thm | False of thm; | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 112 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 113 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 114 | type data = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 115 | skel (* structure of the sum type encoding "program points" *) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 116 | * (int -> typ) (* types of program points *) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 117 | * (term list Inttab.table) (* measures for program points *) | 
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 118 | * (term * term -> thm option) (* which calls form chains? (cached) *) | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 119 | * (term * (term * term) -> cell)(* local descents (cached) *) | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 120 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 121 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 122 | (* Build case expression *) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 123 | fun mk_sumcases (sk, _, _, _, _) T fs = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 124 | mk_tree (fn i => (nth fs i, domain_type (fastype_of (nth fs i)))) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 125 | (fn ((f, fT), (g, gT)) => (SumTree.mk_sumcase fT gT T f g, SumTree.mk_sumT fT gT)) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 126 | sk | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 127 | |> fst | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 128 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 129 | fun mk_sum_skel rel = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 130 | let | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
32683diff
changeset | 131 |     val cs = Function_Lib.dest_binop_list @{const_name Lattices.sup} rel
 | 
| 32602 | 132 |     fun collect_pats (Const (@{const_name Collect}, _) $ Abs (_, _, c)) =
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 133 | let | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 134 |         val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _)
 | 
| 38557 | 135 |           = Term.strip_qnt_body @{const_name Ex} c
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 136 | in cons r o cons l end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 137 | in | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 138 | mk_skel (fold collect_pats cs []) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 139 | end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 140 | |
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 141 | fun prove_chain thy chain_tac (c1, c2) = | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 142 | let | 
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 143 | val goal = | 
| 47433 
07f4bf913230
renamed "rel_comp" to "relcomp" (to be consistent with, e.g., "relpow")
 griff parents: 
46218diff
changeset | 144 |       HOLogic.mk_eq (HOLogic.mk_binop @{const_name Relation.relcomp} (c1, c2),
 | 
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 145 |         Const (@{const_abbrev Set.empty}, fastype_of c1))
 | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 146 |       |> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *)
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 147 | in | 
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 148 | case Function_Lib.try_proof (cterm_of thy goal) chain_tac of | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 149 | Function_Lib.Solved thm => SOME thm | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 150 | | _ => NONE | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 151 | end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 152 | |
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 153 | |
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 154 | fun dest_call' sk (Const (@{const_name Collect}, _) $ Abs (_, _, c)) =
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 155 | let | 
| 38557 | 156 |     val vs = Term.strip_qnt_vars @{const_name Ex} c
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 157 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 158 | (* FIXME: throw error "dest_call" for malformed terms *) | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 159 |     val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ Gam)
 | 
| 38557 | 160 |       = Term.strip_qnt_body @{const_name Ex} c
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 161 | val (p, l') = dest_inj sk l | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 162 | val (q, r') = dest_inj sk r | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 163 | in | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 164 | (vs, p, l', q, r', Gam) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 165 | end | 
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 166 | | dest_call' _ _ = error "dest_call" | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 167 | |
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 168 | fun dest_call (sk, _, _, _, _) = dest_call' sk | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 169 | |
| 34229 | 170 | fun mk_desc thy tac vs Gam l r m1 m2 = | 
| 171 | let | |
| 172 | fun try rel = | |
| 173 | try_proof (cterm_of thy | |
| 46218 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 wenzelm parents: 
45740diff
changeset | 174 | (Logic.list_all (vs, | 
| 34229 | 175 | Logic.mk_implies (HOLogic.mk_Trueprop Gam, | 
| 176 |              HOLogic.mk_Trueprop (Const (rel, @{typ "nat => nat => bool"})
 | |
| 177 | $ (m2 $ r) $ (m1 $ l)))))) tac | |
| 178 | in | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
34974diff
changeset | 179 |     case try @{const_name Orderings.less} of
 | 
| 34229 | 180 | Solved thm => Less thm | 
| 181 | | Stuck thm => | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
34974diff
changeset | 182 |        (case try @{const_name Orderings.less_eq} of
 | 
| 34229 | 183 | Solved thm2 => LessEq (thm2, thm) | 
| 184 | | Stuck thm2 => | |
| 45740 | 185 |           if prems_of thm2 = [HOLogic.Trueprop $ @{term False}]
 | 
| 34229 | 186 | then False thm2 else None (thm2, thm) | 
| 187 | | _ => raise Match) (* FIXME *) | |
| 188 | | _ => raise Match | |
| 189 | end | |
| 190 | ||
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 191 | fun prove_descent thy tac sk (c, (m1, m2)) = | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 192 | let | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 193 | val (vs, _, l, _, r, Gam) = dest_call' sk c | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 194 | in | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 195 | mk_desc thy tac vs Gam l r m1 m2 | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 196 | end | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 197 | |
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 198 | fun create ctxt chain_tac descent_tac T rel = | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 199 | let | 
| 42361 | 200 | val thy = Proof_Context.theory_of ctxt | 
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 201 | val sk = mk_sum_skel rel | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 202 | val Ts = node_types sk T | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 203 | val M = Inttab.make (map_index (apsnd (MeasureFunctions.get_measure_functions ctxt)) Ts) | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 204 | val chain_cache = Cache.create Term2tab.empty Term2tab.lookup Term2tab.update | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 205 | (prove_chain thy chain_tac) | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 206 | val descent_cache = Cache.create Term3tab.empty Term3tab.lookup Term3tab.update | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 207 | (prove_descent thy descent_tac sk) | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 208 | in | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 209 | (sk, nth Ts, M, chain_cache, descent_cache) | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 210 | end | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 211 | |
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 212 | fun get_num_points (sk, _, _, _, _) = | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 213 | let | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 214 | fun num (SLeaf i) = i + 1 | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 215 | | num (SBranch (s, t)) = num t | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 216 | in num sk end | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 217 | |
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 218 | fun get_types (_, T, _, _, _) = T | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 219 | fun get_measures (_, _, M, _, _) = Inttab.lookup_list M | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 220 | |
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 221 | fun get_chain (_, _, _, C, _) c1 c2 = | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 222 | SOME (C (c1, c2)) | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 223 | |
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 224 | fun get_descent (_, _, _, _, D) c m1 m2 = | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 225 | SOME (D (c, (m1, m2))) | 
| 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 226 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 227 | fun CALLS tac i st = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 228 | if Thm.no_prems st then all_tac st | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 229 | else case Thm.term_of (Thm.cprem_of st i) of | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
32683diff
changeset | 230 |     (_ $ (_ $ rel)) => tac (Function_Lib.dest_binop_list @{const_name Lattices.sup} rel, i) st
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 231 | |_ => no_tac st | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 232 | |
| 39924 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 233 | type ttac = data -> int -> tactic | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 234 | |
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 235 | fun TERMINATION ctxt atac tac = | 
| 32602 | 236 |   SUBGOAL (fn (_ $ (Const (@{const_name wf}, wfT) $ rel), i) =>
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 237 | let | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 238 | val (T, _) = HOLogic.dest_prodT (HOLogic.dest_setT (domain_type wfT)) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 239 | in | 
| 39923 
0e1bd289c8ea
use Cache structure instead of passing tables around explicitly
 krauss parents: 
38864diff
changeset | 240 | tac (create ctxt atac atac T rel) i | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 241 | end) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 242 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 243 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 244 | (* A tactic to convert open to closed termination goals *) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 245 | local | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 246 | fun dest_term (t : term) = (* FIXME, cf. Lexicographic order *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 247 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 248 | val (vars, prop) = Function_Lib.dest_all_all t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 249 | val (prems, concl) = Logic.strip_horn prop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 250 | val (lhs, rhs) = concl | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 251 | |> HOLogic.dest_Trueprop | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 252 | |> HOLogic.dest_mem |> fst | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 253 | |> HOLogic.dest_prod | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 254 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 255 | (vars, prems, lhs, rhs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 256 | end | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 257 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 258 | fun mk_pair_compr (T, qs, l, r, conds) = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 259 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 260 | val pT = HOLogic.mk_prodT (T, T) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 261 | val n = length qs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 262 | val peq = HOLogic.eq_const pT $ Bound n $ (HOLogic.pair_const T T $ l $ r) | 
| 45740 | 263 |     val conds' = if null conds then [@{term True}] else conds
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 264 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 265 | HOLogic.Collect_const pT $ | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 266 |     Abs ("uu_", pT,
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 267 | (foldr1 HOLogic.mk_conj (peq :: conds') | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 268 | |> fold_rev (fn v => fn t => HOLogic.exists_const (fastype_of v) $ lambda v t) qs)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 269 | end | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 270 | |
| 52723 
2ebcc81f599c
eliminate duplicate calls when moving to closed form -- graph decomposition chokes on duplicates
 krauss parents: 
47835diff
changeset | 271 | val Un_aci_simps = | 
| 
2ebcc81f599c
eliminate duplicate calls when moving to closed form -- graph decomposition chokes on duplicates
 krauss parents: 
47835diff
changeset | 272 |   map mk_meta_eq @{thms Un_ac Un_absorb}
 | 
| 
2ebcc81f599c
eliminate duplicate calls when moving to closed form -- graph decomposition chokes on duplicates
 krauss parents: 
47835diff
changeset | 273 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 274 | in | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 275 | |
| 30607 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 wenzelm parents: 
30304diff
changeset | 276 | fun wf_union_tac ctxt st = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 277 | let | 
| 42361 | 278 | val thy = Proof_Context.theory_of ctxt | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 279 | val cert = cterm_of (theory_of_thm st) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 280 | val ((_ $ (_ $ rel)) :: ineqs) = prems_of st | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 281 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 282 | fun mk_compr ineq = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 283 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 284 | val (vars, prems, lhs, rhs) = dest_term ineq | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 285 | in | 
| 35625 | 286 | mk_pair_compr (fastype_of lhs, vars, lhs, rhs, map (Object_Logic.atomize_term thy) prems) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 287 | end | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 288 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 289 | val relation = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 290 | if null ineqs | 
| 35402 | 291 |       then Const (@{const_abbrev Set.empty}, fastype_of rel)
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 292 | else map mk_compr ineqs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 293 |         |> foldr1 (HOLogic.mk_binop @{const_name Lattices.sup})
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 294 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 295 | fun solve_membership_tac i = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 296 |       (EVERY' (replicate (i - 2) (rtac @{thm UnI2}))  (* pick the right component of the union *)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 297 |       THEN' (fn j => TRY (rtac @{thm UnI1} j))
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 298 |       THEN' (rtac @{thm CollectI})                    (* unfold comprehension *)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 299 |       THEN' (fn i => REPEAT (rtac @{thm exI} i))      (* Turn existentials into schematic Vars *)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 300 |       THEN' ((rtac @{thm refl})                       (* unification instantiates all Vars *)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 301 |         ORELSE' ((rtac @{thm conjI})
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 302 |           THEN' (rtac @{thm refl})
 | 
| 42793 | 303 | THEN' (blast_tac ctxt))) (* Solve rest of context... not very elegant *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 304 | ) i | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 305 | in | 
| 52723 
2ebcc81f599c
eliminate duplicate calls when moving to closed form -- graph decomposition chokes on duplicates
 krauss parents: 
47835diff
changeset | 306 | (PRIMITIVE (Drule.cterm_instantiate [(cert rel, cert relation)]) | 
| 
2ebcc81f599c
eliminate duplicate calls when moving to closed form -- graph decomposition chokes on duplicates
 krauss parents: 
47835diff
changeset | 307 | THEN ALLGOALS (fn i => if i = 1 then all_tac else solve_membership_tac i) | 
| 52732 | 308 | THEN rewrite_goal_tac Un_aci_simps 1) st (* eliminate duplicates *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 309 | end | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 310 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 311 | end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 312 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 313 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 314 | |
| 34228 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 315 | (*** DEPENDENCY GRAPHS ***) | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 316 | |
| 34228 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 317 | fun mk_dgraph D cs = | 
| 35404 | 318 | Term_Graph.empty | 
| 319 | |> fold (fn c => Term_Graph.new_node (c, ())) cs | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 320 | |> fold_product (fn c1 => fn c2 => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 321 | if is_none (get_chain D c1 c2 |> the_default NONE) | 
| 47835 | 322 | then Term_Graph.add_edge (c2, c1) else I) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 323 | cs cs | 
| 34228 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 324 | |
| 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 325 | fun ucomp_empty_tac T = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 326 |   REPEAT_ALL_NEW (rtac @{thm union_comp_emptyR}
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 327 |     ORELSE' rtac @{thm union_comp_emptyL}
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 328 | ORELSE' SUBGOAL (fn (_ $ (_ $ (_ $ c1 $ c2) $ _), i) => rtac (T c1 c2) i)) | 
| 34228 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 329 | |
| 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 330 | fun regroup_calls_tac cs = CALLS (fn (cs', i) => | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 331 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 332 | val is = map (fn c => find_index (curry op aconv c) cs') cs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 333 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 334 | CONVERSION (Conv.arg_conv (Conv.arg_conv | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 335 | (Function_Lib.regroup_union_conv is))) i | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 336 | end) | 
| 34228 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 337 | |
| 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 338 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 339 | fun solve_trivial_tac D = CALLS (fn ([c], i) => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 340 | (case get_chain D c c of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 341 |      SOME (SOME thm) => rtac @{thm wf_no_loop} i
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 342 | THEN rtac thm i | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 343 | | _ => no_tac) | 
| 34228 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 344 | | _ => no_tac) | 
| 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 345 | |
| 39924 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 346 | fun decompose_tac D = CALLS (fn (cs, i) => | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 347 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 348 | val G = mk_dgraph D cs | 
| 35404 | 349 | val sccs = Term_Graph.strong_conn G | 
| 34228 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 350 | |
| 39924 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 351 | fun split [SCC] i = TRY (solve_trivial_tac D i) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 352 | | split (SCC::rest) i = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 353 | regroup_calls_tac SCC i | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 354 |         THEN rtac @{thm wf_union_compatible} i
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 355 |         THEN rtac @{thm less_by_empty} (i + 2)
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 356 | THEN ucomp_empty_tac (the o the oo get_chain D) (i + 2) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 357 | THEN split rest (i + 1) | 
| 39924 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 358 | THEN TRY (solve_trivial_tac D i) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 359 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 360 | if length sccs > 1 then split sccs i | 
| 39924 
f4d3e70ed3a8
discontinued continuations to simplify control flow; dropped optimization in scnp
 krauss parents: 
39923diff
changeset | 361 | else solve_trivial_tac D i | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
34229diff
changeset | 362 | end) | 
| 34228 
bc0cea4cae52
absorb structures Decompose and Descent into Termination, to simplify further restructuring
 krauss parents: 
33855diff
changeset | 363 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 364 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 365 | end |