author | Christian Urban <urbanc@in.tum.de> |
Thu, 08 Sep 2011 11:21:03 +0100 | |
changeset 3002 | 02d98590454d |
parent 2982 | 4a00077c008f |
child 3045 | d0ad264f8c4f |
child 3068 | f89ee40fbb08 |
permissions | -rw-r--r-- |
2976 | 1 |
(* Nominal Termination |
2 |
Author: Christian Urban |
|
3 |
||
4 |
heavily based on the code of Alexander Krauss |
|
5 |
(code forked on 18 July 2011) |
|
6 |
||
7 |
Redefinition of the termination command |
|
8 |
*) |
|
9 |
||
10 |
signature NOMINAL_FUNCTION_TERMINATION = |
|
11 |
sig |
|
12 |
include NOMINAL_FUNCTION_DATA |
|
13 |
||
14 |
val termination : term option -> local_theory -> Proof.state |
|
15 |
val termination_cmd : string option -> local_theory -> Proof.state |
|
16 |
||
17 |
end |
|
18 |
||
19 |
structure Nominal_Function_Termination : NOMINAL_FUNCTION_TERMINATION = |
|
20 |
struct |
|
21 |
||
22 |
open Function_Lib |
|
23 |
open Function_Common |
|
24 |
open Nominal_Function_Common |
|
25 |
||
26 |
val simp_attribs = map (Attrib.internal o K) |
|
27 |
[Simplifier.simp_add, |
|
28 |
Code.add_default_eqn_attribute, |
|
29 |
Nitpick_Simps.add] |
|
30 |
||
31 |
val eqvt_attrib = Attrib.internal (K Nominal_ThmDecls.eqvt_add) |
|
32 |
||
33 |
fun prepare_termination_proof prep_term raw_term_opt lthy = |
|
34 |
let |
|
35 |
val term_opt = Option.map (prep_term lthy) raw_term_opt |
|
36 |
val info = the (case term_opt of |
|
37 |
SOME t => (import_function_data t lthy |
|
38 |
handle Option.Option => |
|
39 |
error ("Not a function: " ^ quote (Syntax.string_of_term lthy t))) |
|
40 |
| NONE => (import_last_function lthy handle Option.Option => error "Not a function")) |
|
41 |
||
42 |
val { termination, fs, R, add_simps, case_names, psimps, |
|
43 |
pinducts, defname, eqvts, ...} = info |
|
44 |
val domT = domain_type (fastype_of R) |
|
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
45 |
|
2976 | 46 |
val goal = HOLogic.mk_Trueprop |
47 |
(HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT))) |
|
48 |
fun afterqed [[totality]] lthy = |
|
49 |
let |
|
50 |
val totality = Thm.close_derivation totality |
|
51 |
val remove_domain_condition = |
|
52 |
full_simplify (HOL_basic_ss addsimps [totality, @{thm True_implies_equals}]) |
|
53 |
val tsimps = map remove_domain_condition psimps |
|
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
54 |
val tinducts = map remove_domain_condition pinducts |
2976 | 55 |
val teqvts = map remove_domain_condition eqvts |
2981 | 56 |
|
2976 | 57 |
fun qualify n = Binding.name n |
58 |
|> Binding.qualify true defname |
|
59 |
in |
|
60 |
lthy |
|
61 |
|> add_simps I "simps" I simp_attribs tsimps |
|
62 |
||>> Local_Theory.note ((qualify "eqvt", [eqvt_attrib]), teqvts) |
|
63 |
||>> Local_Theory.note |
|
64 |
((qualify "induct", |
|
65 |
[Attrib.internal (K (Rule_Cases.case_names case_names))]), |
|
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
66 |
tinducts) |
2976 | 67 |
|-> (fn ((simps, (_, eqvts)), (_, inducts)) => fn lthy => |
68 |
let val info' = { is_partial=false, defname=defname, add_simps=add_simps, |
|
69 |
case_names=case_names, fs=fs, R=R, psimps=psimps, pinducts=pinducts, |
|
2981 | 70 |
simps=SOME simps, inducts=SOME inducts, termination=termination, eqvts=teqvts } |
2976 | 71 |
in |
72 |
(info', |
|
73 |
lthy |
|
74 |
|> Local_Theory.declaration false (add_function_data o morph_function_data info') |
|
75 |
|> Spec_Rules.add Spec_Rules.Equational (fs, tsimps)) |
|
76 |
end) |
|
77 |
end |
|
78 |
in |
|
79 |
(goal, afterqed, termination) |
|
80 |
end |
|
81 |
||
82 |
fun gen_termination prep_term raw_term_opt lthy = |
|
83 |
let |
|
84 |
val (goal, afterqed, termination) = prepare_termination_proof prep_term raw_term_opt lthy |
|
85 |
in |
|
86 |
lthy |
|
87 |
|> Proof_Context.note_thmss "" |
|
88 |
[((Binding.empty, [Context_Rules.rule_del]), [([allI], [])])] |> snd |
|
89 |
|> Proof_Context.note_thmss "" |
|
90 |
[((Binding.empty, [Context_Rules.intro_bang (SOME 1)]), [([allI], [])])] |> snd |
|
91 |
|> Proof_Context.note_thmss "" |
|
92 |
[((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]), |
|
93 |
[([Goal.norm_result termination], [])])] |> snd |
|
94 |
|> Proof.theorem NONE (snd oo afterqed) [[(goal, [])]] |
|
95 |
end |
|
96 |
||
97 |
val termination = gen_termination Syntax.check_term |
|
98 |
val termination_cmd = gen_termination Syntax.read_term |
|
99 |
||
100 |
(* outer syntax *) |
|
101 |
||
102 |
val option_parser = |
|
103 |
(Scan.optional (Parse.$$$ "(" |-- Parse.!!! |
|
104 |
(Parse.reserved "eqvt" >> K true) --| Parse.$$$ ")") false) |
|
105 |
||
106 |
val _ = |
|
107 |
Outer_Syntax.local_theory_to_proof "termination" "prove termination of a recursive function" |
|
108 |
Keyword.thy_goal |
|
109 |
(option_parser -- Scan.option Parse.term >> |
|
110 |
(fn (is_eqvt, trm) => |
|
111 |
if is_eqvt then termination_cmd trm else Function.termination_cmd trm)) |
|
112 |
||
113 |
end |