| author | haftmann |
| Fri, 13 Oct 2006 16:52:49 +0200 | |
| changeset 21014 | 3b0c2641f740 |
| parent 21000 | 54c42dfbcafa |
| child 21051 | c49467a9c1e1 |
| 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/fundef_package.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 |
|
|
20363
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20338
diff
changeset
|
5 |
A package for general recursive function definitions. |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
6 |
Isar commands. |
|
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 |
*) |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
9 |
|
|
20363
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20338
diff
changeset
|
10 |
signature FUNDEF_PACKAGE = |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
11 |
sig |
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
12 |
val add_fundef : (string * string option * mixfix) list |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
13 |
-> ((bstring * Attrib.src list) * string list) list list |
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
14 |
-> FundefCommon.fundef_config |
|
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
15 |
-> local_theory |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
16 |
-> Proof.state |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
17 |
|
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
18 |
val add_fundef_i: (string * typ option * mixfix) list |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
19 |
-> ((bstring * Attrib.src list) * term list) list list |
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
20 |
-> bool |
|
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
21 |
-> local_theory |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
22 |
-> Proof.state |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
23 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
24 |
val cong_add: attribute |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
25 |
val cong_del: attribute |
|
20363
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20338
diff
changeset
|
26 |
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
27 |
val setup : theory -> theory |
|
19770
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
19617
diff
changeset
|
28 |
val get_congs : theory -> thm list |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
29 |
end |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
30 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
31 |
|
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
32 |
structure FundefPackage = |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
33 |
struct |
|
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 |
open FundefCommon |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
36 |
|
|
20654
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents:
20638
diff
changeset
|
37 |
(*fn (ctxt,s) => Variable.importT_terms (fst (ProofContext.read_termTs ctxt (K false) (K NONE) (K NONE) [] [s])) ctxt;*) |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
38 |
|
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
39 |
fun burrow_snd f ps = (* ('a list -> 'b list) -> ('c * 'a) list -> ('c * 'b) list *)
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
40 |
let val (xs, ys) = split_list ps |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
41 |
in xs ~~ f ys end |
|
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
19938
diff
changeset
|
42 |
|
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
43 |
fun restore_spec_structure reps spec = |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
44 |
(burrow o burrow_snd o burrow o K) reps spec |
|
20363
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20338
diff
changeset
|
45 |
|
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
46 |
fun add_simps label moreatts mutual_info fixes psimps spec lthy = |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
47 |
let |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
48 |
val fnames = map (fst o fst) fixes |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
49 |
val psimps_by_f = FundefMutual.sort_by_function mutual_info fnames psimps |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
50 |
|
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
51 |
fun add_for_f fname psimps = |
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
52 |
LocalTheory.note ((NameSpace.append fname label, Attrib.internal Simplifier.simp_add :: moreatts), psimps) #> snd |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
53 |
in |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
54 |
lthy |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
55 |
|> fold_map (fold_map LocalTheory.note) (restore_spec_structure psimps spec) |> snd |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
56 |
|> fold2 add_for_f fnames psimps_by_f |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
57 |
end |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
58 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
59 |
|
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
60 |
fun fundef_afterqed fixes spec mutual_info defname data [[result]] lthy = |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
61 |
let |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
62 |
val fundef_data = FundefMutual.mk_partial_rules_mutual lthy mutual_info data result |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
63 |
val FundefMResult {psimps, subset_pinducts, simple_pinducts, termination, domintros, cases, ...} = fundef_data
|
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
64 |
val qualify = NameSpace.append defname; |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
65 |
in |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
66 |
lthy |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
67 |
|> add_simps "psimps" [] mutual_info fixes psimps spec |
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
68 |
|> LocalTheory.note ((qualify "domintros", []), domintros) |> snd |
|
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
69 |
|> LocalTheory.note ((qualify "termination", []), [termination]) |> snd |
|
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
70 |
|> LocalTheory.note ((qualify "cases", []), [cases]) |> snd |
|
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
71 |
|> LocalTheory.note ((qualify "pinduct", [Attrib.internal (InductAttrib.induct_set "")]), simple_pinducts) |> snd |
|
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
72 |
|> LocalTheory.declaration (add_fundef_data defname (fundef_data, mutual_info, (fixes,spec))) |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
73 |
end (* FIXME: Add cases for induct and cases thm *) |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
74 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
75 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
76 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
77 |
fun prep_with_flags prep fixspec eqnss_flags global_flag lthy = |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
78 |
let |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
79 |
val eqnss = map (map (apsnd (map fst))) eqnss_flags |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
80 |
val flags = map (map (map (fn (_, f) => global_flag orelse f) o snd)) eqnss_flags |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
81 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
82 |
val ((fixes, _), ctxt') = prep fixspec [] lthy |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
83 |
val spec = map (fn eqns => snd (fst (prep [] eqns ctxt'))) eqnss |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
84 |
|> map (map (apsnd (map (fn t => fold_rev (mk_forall o Free) (frees_in_term ctxt' t) t)))) (* Add quantifiers *) |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
85 |
|> map2 (map2 (fn fs => fn (r, thms) => (r, fs ~~ thms))) flags |
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
86 |
|> (burrow o burrow_snd o burrow) |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
87 |
(FundefSplit.split_some_equations lthy) |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
88 |
|> map (map (apsnd flat)) |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
89 |
in |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
90 |
((fixes, spec), ctxt') |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
91 |
end |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
92 |
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
93 |
|
|
20654
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents:
20638
diff
changeset
|
94 |
fun gen_add_fundef prep_spec fixspec eqnss_flags config lthy = |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
95 |
let |
|
20654
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents:
20638
diff
changeset
|
96 |
val FundefConfig {sequential, default, ...} = config
|
|
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents:
20638
diff
changeset
|
97 |
|
|
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents:
20638
diff
changeset
|
98 |
val ((fixes, spec), ctxt') = prep_with_flags prep_spec fixspec eqnss_flags sequential lthy |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
99 |
val t_eqns = spec |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
100 |
|> flat |> map snd |> flat (* flatten external structure *) |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
101 |
|
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
102 |
val ((mutual_info, name, prep_result as Prep {goal, goalI, ...}), lthy) =
|
|
20654
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents:
20638
diff
changeset
|
103 |
FundefMutual.prepare_fundef_mutual fixes t_eqns default lthy |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
104 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
105 |
val afterqed = fundef_afterqed fixes spec mutual_info name prep_result |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
106 |
in |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
107 |
lthy |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
108 |
|> Proof.theorem_i PureThy.internalK NONE afterqed NONE ("", []) [(("", []), [(goal, [])])]
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
109 |
|> Proof.refine (Method.primitive_text (fn _ => goalI)) |> Seq.hd |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
110 |
end |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
111 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
112 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
113 |
fun total_termination_afterqed defname data [[totality]] lthy = |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
114 |
let |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
115 |
val (FundefMResult {psimps, simple_pinducts, ... }, mutual_info, (fixes, stmts)) = data
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
116 |
|
|
20363
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20338
diff
changeset
|
117 |
val remove_domain_condition = full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals]) |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
118 |
|
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
119 |
val tsimps = map remove_domain_condition psimps |
|
20363
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20338
diff
changeset
|
120 |
val tinduct = map remove_domain_condition simple_pinducts |
|
19770
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
19617
diff
changeset
|
121 |
|
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
122 |
(* FIXME: How to generate code from (possibly) local contexts |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
123 |
val has_guards = exists ((fn (Const ("Trueprop", _) $ _) => false | _ => true) o prop_of) tsimps
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
124 |
val allatts = if has_guards then [] else [Attrib.internal (RecfunCodegen.add NONE)] |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
125 |
*) |
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
126 |
val qualify = NameSpace.append defname; |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
127 |
in |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
128 |
lthy |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
129 |
|> add_simps "simps" [] mutual_info fixes tsimps stmts |
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
130 |
|> LocalTheory.note ((qualify "induct", []), tinduct) |> snd |
|
20363
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20338
diff
changeset
|
131 |
end |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
132 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
133 |
|
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
134 |
fun fundef_setup_termination_proof name_opt lthy = |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
135 |
let |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
136 |
val name = the_default (get_last_fundef (Context.Proof lthy)) name_opt |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
137 |
val data = the (get_fundef_data name (Context.Proof lthy)) |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
138 |
handle Option.Option => raise ERROR ("No such function definition: " ^ name)
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
139 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
140 |
val (res as FundefMResult {termination, ...}, _, _) = data
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
141 |
val goal = FundefTermination.mk_total_termination_goal data |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
142 |
in |
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
143 |
lthy |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
144 |
|> ProofContext.note_thmss_i [(("termination",
|
|
20877
368b997ad67e
removed with_local_path -- LocalTheory.note admits qualified names;
wenzelm
parents:
20654
diff
changeset
|
145 |
[ContextRules.intro_query NONE]), [(ProofContext.standard lthy [termination], [])])] |> snd |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
146 |
|> Proof.theorem_i PureThy.internalK NONE |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
147 |
(total_termination_afterqed name data) NONE ("", [])
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
148 |
[(("", []), [(goal, [])])]
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
149 |
end |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
150 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
151 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
152 |
val add_fundef = gen_add_fundef Specification.read_specification |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
153 |
val add_fundef_i = gen_add_fundef Specification.cert_specification |
|
19611
da2a0014c461
Function Package: Quick-and-dirty-fixed strange "Proved a different theorem bug"
krauss
parents:
19585
diff
changeset
|
154 |
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
155 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
156 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
157 |
(* congruence rules *) |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
158 |
|
| 19617 | 159 |
val cong_add = Thm.declaration_attribute (map_fundef_congs o Drule.add_rule o safe_mk_meta_eq); |
160 |
val cong_del = Thm.declaration_attribute (map_fundef_congs o Drule.del_rule o safe_mk_meta_eq); |
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
161 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
162 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
163 |
(* setup *) |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
164 |
|
|
20363
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20338
diff
changeset
|
165 |
val setup = FundefData.init #> FundefCongs.init |
|
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20338
diff
changeset
|
166 |
#> Attrib.add_attributes |
|
f34c5dbe74d5
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20338
diff
changeset
|
167 |
[("fundef_cong", Attrib.add_del_args cong_add cong_del, "declaration of congruence rule for function definitions")]
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
168 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
169 |
|
|
19770
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
19617
diff
changeset
|
170 |
val get_congs = FundefCommon.get_fundef_congs o Context.Theory |
|
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
19617
diff
changeset
|
171 |
|
|
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
19617
diff
changeset
|
172 |
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
173 |
(* outer syntax *) |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
174 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
175 |
local structure P = OuterParse and K = OuterKeyword in |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
176 |
|
|
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
19938
diff
changeset
|
177 |
|
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
19938
diff
changeset
|
178 |
|
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
179 |
fun or_list1 s = P.enum1 "|" s |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
180 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
181 |
val opt_sequential = Scan.optional ((P.$$$ "(" |-- P.$$$ "sequential" --| P.$$$ ")") >> K true) false
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
182 |
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
183 |
val otherwise = P.$$$ "(" |-- P.$$$ "otherwise" --| P.$$$ ")"
|
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
184 |
val statement_ow = P.and_list1 (P.opt_thm_name ":" -- Scan.repeat1 (P.prop -- Scan.optional (otherwise >> K true) false)) |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
185 |
val statements_ow = or_list1 statement_ow |
|
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
19938
diff
changeset
|
186 |
|
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
19938
diff
changeset
|
187 |
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
188 |
val functionP = |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
189 |
OuterSyntax.command "function" "define general recursive functions" K.thy_goal |
|
20654
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents:
20638
diff
changeset
|
190 |
((config_parser -- P.opt_locale_target -- P.fixes --| P.$$$ "where" -- statements_ow) |
|
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents:
20638
diff
changeset
|
191 |
>> (fn (((config, target), fixes), statements) => |
| 21000 | 192 |
Toplevel.print o |
193 |
Toplevel.local_theory_to_proof target (add_fundef fixes statements config))); |
|
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20363
diff
changeset
|
194 |
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
195 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
196 |
val terminationP = |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
197 |
OuterSyntax.command "termination" "prove termination of a recursive function" K.thy_goal |
| 21000 | 198 |
(P.opt_locale_target -- Scan.option P.name |
199 |
>> (fn (target, name) => |
|
200 |
Toplevel.print o |
|
201 |
Toplevel.local_theory_to_proof target (fundef_setup_termination_proof name))); |
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
202 |
|
|
20338
ecdfc96cf4d0
Added Keywords: "otherwise" and "sequential", needed for function package's
krauss
parents:
20320
diff
changeset
|
203 |
|
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
204 |
val _ = OuterSyntax.add_parsers [functionP]; |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
205 |
val _ = OuterSyntax.add_parsers [terminationP]; |
|
20654
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents:
20638
diff
changeset
|
206 |
val _ = OuterSyntax.add_keywords ["sequential", "otherwise"]; |
|
19564
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
207 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
208 |
end; |
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
209 |
|
|
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff
changeset
|
210 |
|
| 19585 | 211 |
end |