| author | wenzelm | 
| Thu, 05 Nov 2009 15:54:14 +0100 | |
| changeset 33448 | 7f716a975ada | 
| parent 33099 | b8cdd3d73022 | 
| permissions | -rw-r--r-- | 
| 31775 | 1 | (* Title: HOL/Tools/Function/descent.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 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 4 | Descent proofs for termination | 
| 
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 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 8 | signature DESCENT = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 9 | sig | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 10 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 11 | val derive_diag : Proof.context -> tactic -> (Termination.data -> int -> tactic) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 12 | -> Termination.data -> int -> tactic | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 13 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 14 | val derive_all : Proof.context -> tactic -> (Termination.data -> int -> tactic) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 15 | -> Termination.data -> int -> tactic | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 16 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 17 | end | 
| 
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 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 20 | structure Descent : DESCENT = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 21 | struct | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 22 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 23 | fun gen_descent diag ctxt tac cont D = Termination.CALLS (fn (cs, i) => | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 24 | let | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 25 | val thy = ProofContext.theory_of ctxt | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 26 | val measures_of = Termination.get_measures D | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 27 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 28 | fun derive c D = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 29 | let | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 30 | val (_, p, _, q, _, _) = Termination.dest_call D c | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 31 | in | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 32 | if diag andalso p = q | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 33 | then fold (fn m => Termination.derive_descent thy tac c m m) (measures_of p) D | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 34 | else fold_product (Termination.derive_descent thy tac c) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 35 | (measures_of p) (measures_of q) D | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 36 | end | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 37 | in | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
31775diff
changeset | 38 | cont (Function_Common.PROFILE "deriving descents" (fold derive cs) D) i | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 39 | end) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 40 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 41 | val derive_diag = gen_descent true | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 42 | val derive_all = gen_descent false | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 43 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: diff
changeset | 44 | end |