author | haftmann |
Tue, 08 Jun 2010 16:37:19 +0200 | |
changeset 37387 | 3581483cca6c |
parent 36636 | 7dded80a953f |
child 40057 | b237f757b215 |
permissions | -rw-r--r-- |
33100 | 1 |
(* Title: HOL/Tools/Function/relation.ML |
2 |
Author: Alexander Krauss, TU Muenchen |
|
3 |
||
4 |
A package for general recursive function definitions. |
|
5 |
Method "relation" to commence a termination proof using a user-specified relation. |
|
6 |
*) |
|
7 |
||
8 |
signature FUNCTION_RELATION = |
|
9 |
sig |
|
10 |
val relation_tac: Proof.context -> term -> int -> tactic |
|
11 |
val setup: theory -> theory |
|
12 |
end |
|
13 |
||
14 |
structure Function_Relation : FUNCTION_RELATION = |
|
15 |
struct |
|
16 |
||
36636 | 17 |
fun inst_state_tac ctxt rel st = |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33100
diff
changeset
|
18 |
let |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33100
diff
changeset
|
19 |
val cert = Thm.cterm_of (ProofContext.theory_of ctxt) |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33100
diff
changeset
|
20 |
val rel' = cert (singleton (Variable.polymorphic ctxt) rel) |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33100
diff
changeset
|
21 |
val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st |
36636 | 22 |
in case Term.add_vars (prop_of st') [] of |
23 |
[v] => |
|
24 |
PRIMITIVE (Drule.cterm_instantiate [(cert (Var v), rel')]) st' |
|
25 |
| _ => Seq.empty |
|
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33100
diff
changeset
|
26 |
end |
33100 | 27 |
|
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33100
diff
changeset
|
28 |
fun relation_tac ctxt rel i = |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33100
diff
changeset
|
29 |
TRY (Function_Common.apply_termination_rule ctxt i) |
36636 | 30 |
THEN inst_state_tac ctxt rel |
33100 | 31 |
|
32 |
val setup = |
|
33 |
Method.setup @{binding relation} |
|
34 |
(Args.term >> (fn rel => fn ctxt => SIMPLE_METHOD' (relation_tac ctxt rel))) |
|
35 |
"proves termination using a user-specified wellfounded relation" |
|
36 |
||
37 |
end |