| author | urbanc |
| Thu, 22 Mar 2007 11:17:32 +0100 | |
| changeset 22495 | c54748fd1f43 |
| parent 21879 | a3efbae45735 |
| child 22496 | 1f428200fd9c |
| permissions | -rw-r--r-- |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
1 |
(* Title: HOL/Tools/function_package/auto_term.ML |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
2 |
ID: $Id$ |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
3 |
Author: Alexander Krauss, TU Muenchen |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
4 |
|
| 20873 | 5 |
A package for general recursive function definitions. |
|
21319
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
6 |
Method "relation" to commence a termination proof using a user-specified relation. |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
7 |
*) |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
8 |
|
|
21319
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
9 |
signature FUNDEF_RELATION = |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
10 |
sig |
| 21588 | 11 |
val relation_tac: Proof.context -> term -> int -> tactic |
| 20873 | 12 |
val setup: theory -> theory |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
13 |
end |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
14 |
|
|
21319
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
15 |
structure FundefRelation : FUNDEF_RELATION = |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
16 |
struct |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
17 |
|
| 21588 | 18 |
fun relation_tac ctxt rel = |
|
21319
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
19 |
let |
|
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
20 |
val cert = Thm.cterm_of (ProofContext.theory_of ctxt) |
| 21588 | 21 |
|
|
21319
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
22 |
val rel' = cert (singleton (Variable.polymorphic ctxt) rel) |
|
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
23 |
val rule = FundefCommon.get_termination_rule ctxt |> the |
|
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
24 |
|> Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) |
| 21588 | 25 |
|
|
21319
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
26 |
val prem = #1 (Logic.dest_implies (Thm.prop_of rule)) |
|
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
27 |
val Rvar = cert (Var (the_single (Term.add_vars prem []))) |
| 21588 | 28 |
in rtac (Drule.cterm_instantiate [(Rvar, rel')] rule) end |
29 |
||
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
30 |
|
| 20295 | 31 |
val setup = Method.add_methods |
| 21879 | 32 |
[("relation", (Method.SIMPLE_METHOD' o (fn (rel, ctxt) => relation_tac ctxt rel)) oo (Method.syntax Args.term),
|
|
21319
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
krauss
parents:
21104
diff
changeset
|
33 |
"proves termination using a user-specified wellfounded relation")] |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
34 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
35 |
end |