author | nipkow |
Tue, 17 Jun 2025 06:29:55 +0200 | |
changeset 82732 | 71574900b6ba |
parent 77879 | dd222e2af01a |
permissions | -rw-r--r-- |
33100 | 1 |
(* Title: HOL/Tools/Function/relation.ML |
2 |
Author: Alexander Krauss, TU Muenchen |
|
3 |
||
47701 | 4 |
Tactics to start a termination proof using a user-specified relation. |
33100 | 5 |
*) |
6 |
||
7 |
signature FUNCTION_RELATION = |
|
8 |
sig |
|
40445
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
krauss
parents:
40057
diff
changeset
|
9 |
val relation_tac: Proof.context -> (typ -> term) -> int -> tactic |
47701 | 10 |
val relation_infer_tac: Proof.context -> term -> int -> tactic |
33100 | 11 |
end |
12 |
||
59968 | 13 |
structure Function_Relation: FUNCTION_RELATION = |
33100 | 14 |
struct |
15 |
||
40445
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
krauss
parents:
40057
diff
changeset
|
16 |
(* tactic version *) |
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
krauss
parents:
40057
diff
changeset
|
17 |
|
59968 | 18 |
fun inst_state_tac ctxt inst = |
19 |
SUBGOAL (fn (goal, _) => |
|
20 |
(case Term.add_vars goal [] of |
|
74282 | 21 |
[v as (_, T)] => |
77879 | 22 |
PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make1 (v, Thm.cterm_of ctxt (inst T)))) |
59968 | 23 |
| _ => no_tac)); |
40445
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
krauss
parents:
40057
diff
changeset
|
24 |
|
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
krauss
parents:
40057
diff
changeset
|
25 |
fun relation_tac ctxt rel i = |
59159 | 26 |
TRY (Function_Common.termination_rule_tac ctxt i) |
59968 | 27 |
THEN inst_state_tac ctxt rel i; |
40445
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
krauss
parents:
40057
diff
changeset
|
28 |
|
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
krauss
parents:
40057
diff
changeset
|
29 |
|
47701 | 30 |
(* version with type inference *) |
40445
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
krauss
parents:
40057
diff
changeset
|
31 |
|
59968 | 32 |
fun inst_state_infer_tac ctxt rel = |
33 |
SUBGOAL (fn (goal, _) => |
|
34 |
(case Term.add_vars goal [] of |
|
35 |
[v as (_, T)] => |
|
36 |
let |
|
37 |
val rel' = |
|
38 |
singleton (Variable.polymorphic ctxt) rel |
|
39 |
|> map_types Type_Infer.paramify_vars |
|
40 |
|> Type.constraint T |
|
41 |
|> Syntax.check_term ctxt; |
|
77879 | 42 |
in PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make1 (v, Thm.cterm_of ctxt rel'))) end |
59968 | 43 |
| _ => no_tac)); |
33100 | 44 |
|
47701 | 45 |
fun relation_infer_tac ctxt rel i = |
59159 | 46 |
TRY (Function_Common.termination_rule_tac ctxt i) |
59968 | 47 |
THEN inst_state_infer_tac ctxt rel i; |
33100 | 48 |
|
49 |
end |