| author | wenzelm | 
| Mon, 16 Mar 2015 16:59:59 +0100 | |
| changeset 59720 | f893472fff31 | 
| parent 59627 | bb1e4a35d506 | 
| child 59968 | d69dc7a133e7 | 
| 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: 
40057diff
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 | ||
| 13 | structure Function_Relation : FUNCTION_RELATION = | |
| 14 | struct | |
| 15 | ||
| 40445 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 16 | (* tactic version *) | 
| 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 17 | |
| 59627 | 18 | fun inst_state_tac ctxt inst st = | 
| 59618 | 19 | (case Term.add_vars (Thm.prop_of st) [] of (* FIXME tactic should not inspect main conclusion *) | 
| 40445 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 20 | [v as (_, T)] => | 
| 59627 | 21 | PRIMITIVE (Thm.instantiate ([], [(Thm.cterm_of ctxt (Var v), Thm.cterm_of ctxt (inst T))])) st | 
| 59582 | 22 | | _ => Seq.empty); | 
| 40445 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 23 | |
| 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 24 | fun relation_tac ctxt rel i = | 
| 59159 | 25 | TRY (Function_Common.termination_rule_tac ctxt i) | 
| 59627 | 26 | THEN inst_state_tac ctxt rel; | 
| 40445 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 27 | |
| 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 28 | |
| 47701 | 29 | (* version with type inference *) | 
| 40445 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 30 | |
| 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 31 | fun inst_state_infer_tac ctxt rel st = | 
| 59618 | 32 | (case Term.add_vars (Thm.prop_of st) [] of (* FIXME tactic should not inspect main conclusion *) | 
| 40057 
b237f757b215
relation method: re-check given term with type constraints to avoid unspecific failure if ill-typed -- keep old behaviour for tactic version
 krauss parents: 
36636diff
changeset | 33 | [v as (_, T)] => | 
| 
b237f757b215
relation method: re-check given term with type constraints to avoid unspecific failure if ill-typed -- keep old behaviour for tactic version
 krauss parents: 
36636diff
changeset | 34 | let | 
| 40445 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 35 | val rel' = singleton (Variable.polymorphic ctxt) rel | 
| 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 36 | |> map_types Type_Infer.paramify_vars | 
| 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 37 | |> Type.constraint T | 
| 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 38 | |> Syntax.check_term ctxt | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 39 | |> Thm.cterm_of ctxt; | 
| 40445 
65bd37d7d501
removed type-inference-like behaviour from relation_tac completely; tuned
 krauss parents: 
40057diff
changeset | 40 | in | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 41 | PRIMITIVE (Thm.instantiate ([], [(Thm.cterm_of ctxt (Var v), rel')])) st | 
| 40057 
b237f757b215
relation method: re-check given term with type constraints to avoid unspecific failure if ill-typed -- keep old behaviour for tactic version
 krauss parents: 
36636diff
changeset | 42 | end | 
| 59582 | 43 | | _ => Seq.empty); | 
| 33100 | 44 | |
| 47701 | 45 | fun relation_infer_tac ctxt rel i = | 
| 59159 | 46 | TRY (Function_Common.termination_rule_tac ctxt i) | 
| 47701 | 47 | THEN inst_state_infer_tac ctxt rel; | 
| 33100 | 48 | |
| 49 | end |