| author | krauss | 
| Tue, 26 Oct 2010 13:19:31 +0200 | |
| changeset 40180 | c3ef007115a0 | 
| parent 40057 | b237f757b215 | 
| child 40445 | 65bd37d7d501 | 
| 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 | ||
| 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 | 17 | fun gen_inst_state_tac prep ctxt rel st = | 
| 
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 | 18 | case Term.add_vars (prop_of st) [] of | 
| 
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 | 19 | [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 | 20 | let | 
| 
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 | 21 | val cert = Thm.cterm_of (ProofContext.theory_of ctxt) | 
| 
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 | 22 | val rel' = prep T (singleton (Variable.polymorphic ctxt) rel) | 
| 
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 | 23 | |> cert | 
| 
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 | 24 | val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st (*FIXME??*) | 
| 
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 | 25 | in | 
| 
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 | 26 | PRIMITIVE (Drule.cterm_instantiate [(cert (Var v), rel')]) st' | 
| 
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 | 27 | end | 
| 
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 | 28 | | _ => Seq.empty; | 
| 33100 | 29 | |
| 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 | 30 | fun gen_relation_tac prep ctxt rel i = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33100diff
changeset | 31 | TRY (Function_Common.apply_termination_rule ctxt i) | 
| 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 | 32 | THEN gen_inst_state_tac prep ctxt rel | 
| 
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 | |
| 
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 | val relation_tac = gen_relation_tac (K I) | 
| 
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 | 35 | |
| 
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 | 36 | fun relation_meth rel ctxt = SIMPLE_METHOD' | 
| 
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 | 37 | (gen_relation_tac (fn T => map_types Type_Infer.paramify_vars | 
| 
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 | 38 | #> Type.constraint T #> Syntax.check_term ctxt) | 
| 
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 | 39 | ctxt rel) | 
| 33100 | 40 | |
| 41 | val setup = | |
| 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 |   Method.setup @{binding relation} (Args.term >> relation_meth)
 | 
| 33100 | 43 | "proves termination using a user-specified wellfounded relation" | 
| 44 | ||
| 45 | end |