author | wenzelm |
Fri, 12 Aug 2016 13:34:59 +0200 | |
changeset 63673 | 2314e99c18a7 |
parent 61268 | abe08fb15a12 |
child 67149 | e61557884799 |
permissions | -rw-r--r-- |
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
30364
diff
changeset
|
1 |
(* Title: HOL/Tools/choice_specification.ML |
14115 | 2 |
Author: Sebastian Skalberg, TU Muenchen |
3 |
||
4 |
Package for defining constants by specification. |
|
5 |
*) |
|
6 |
||
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
30364
diff
changeset
|
7 |
signature CHOICE_SPECIFICATION = |
14115 | 8 |
sig |
35807
e4d1b5cbd429
added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents:
35625
diff
changeset
|
9 |
val close_form : term -> term |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
10 |
val add_specification: (bstring * xstring * bool) list -> theory * thm -> theory * thm |
14115 | 11 |
end |
12 |
||
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
30364
diff
changeset
|
13 |
structure Choice_Specification: CHOICE_SPECIFICATION = |
14115 | 14 |
struct |
15 |
||
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
16 |
local |
46974
7ca3608146d8
eliminated odd 'finalconsts' / Theory.add_finals;
wenzelm
parents:
46961
diff
changeset
|
17 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
18 |
fun mk_definitional [] arg = arg |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
19 |
| mk_definitional ((thname, cname, covld) :: cos) (thy, thm) = |
59582 | 20 |
(case HOLogic.dest_Trueprop (Thm.concl_of thm) of |
21 |
Const (@{const_name Ex}, _) $ P => |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
22 |
let |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
23 |
val ctype = domain_type (type_of P) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
24 |
val cname_full = Sign.intern_const thy cname |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
25 |
val cdefname = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
26 |
if thname = "" |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
27 |
then Thm.def_name (Long_Name.base_name cname) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
28 |
else thname |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
29 |
val def_eq = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
30 |
Logic.mk_equals (Const (cname_full, ctype), HOLogic.choice_const ctype $ P) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
31 |
val (thms, thy') = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
32 |
Global_Theory.add_defs covld [((Binding.name cdefname, def_eq), [])] thy |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
33 |
val thm' = [thm,hd thms] MRS @{thm exE_some} |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
34 |
in |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
35 |
mk_definitional cos (thy',thm') |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
36 |
end |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
37 |
| _ => raise THM ("Bad specification theorem", 0, [thm])) |
14115 | 38 |
|
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
39 |
in |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
40 |
|
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
41 |
fun add_specification cos = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
42 |
mk_definitional cos #> apsnd Drule.export_without_context |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
43 |
|
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
44 |
end |
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
45 |
|
18729 | 46 |
|
14115 | 47 |
(* Collect all intances of constants in term *) |
48 |
||
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
49 |
fun collect_consts (t $ u, tms) = collect_consts (u, collect_consts (t, tms)) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
50 |
| collect_consts (Abs (_, _, t), tms) = collect_consts (t, tms) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
51 |
| collect_consts (tm as Const _, tms) = insert (op aconv) tm tms |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
52 |
| collect_consts (_, tms) = tms |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
53 |
|
14115 | 54 |
|
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35807
diff
changeset
|
55 |
(* Complementing Type.varify_global... *) |
14115 | 56 |
|
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35807
diff
changeset
|
57 |
fun unvarify_global t fmap = |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
58 |
let |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
59 |
val fmap' = map Library.swap fmap |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
60 |
fun unthaw (f as (a, S)) = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
61 |
(case AList.lookup (op =) fmap' a of |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
62 |
NONE => TVar f |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
63 |
| SOME (b, _) => TFree (b, S)) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
64 |
in map_types (map_type_tvar unthaw) t end |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
65 |
|
14115 | 66 |
|
67 |
(* The syntactic meddling needed to setup add_specification for work *) |
|
68 |
||
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
69 |
fun close_form t = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
70 |
fold_rev (fn (s, T) => fn t => HOLogic.mk_all (s, T, t)) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
71 |
(map dest_Free (Misc_Legacy.term_frees t)) t |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
50239
diff
changeset
|
72 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
73 |
fun zip3 [] [] [] = [] |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
74 |
| zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
75 |
| zip3 _ _ _ = raise Fail "Choice_Specification.process_spec internal error" |
14167 | 76 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
77 |
fun myfoldr f [x] = x |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
78 |
| myfoldr f (x::xs) = f (x,myfoldr f xs) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
79 |
| myfoldr f [] = raise Fail "Choice_Specification.process_spec internal error" |
14164
8c3fab596219
Allowed for splitting the specification over several lemmas.
skalberg
parents:
14122
diff
changeset
|
80 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
81 |
fun process_spec cos alt_props thy = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
82 |
let |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
83 |
val ctxt = Proof_Context.init_global thy |
14166 | 84 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
85 |
val rew_imps = alt_props |> |
59642 | 86 |
map (Object_Logic.atomize ctxt o Thm.cterm_of ctxt o Syntax.read_prop_global thy o snd) |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
87 |
val props' = rew_imps |> |
59582 | 88 |
map (HOLogic.dest_Trueprop o Thm.term_of o snd o Thm.dest_equals o Thm.cprop_of) |
14166 | 89 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
90 |
fun proc_single prop = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
91 |
let |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
92 |
val frees = Misc_Legacy.term_frees prop |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
93 |
val _ = forall (fn v => Sign.of_sort thy (type_of v, @{sort type})) frees |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
94 |
orelse error "Specificaton: Only free variables of sort 'type' allowed" |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
95 |
val prop_closed = close_form prop |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
96 |
in (prop_closed, frees) end |
14164
8c3fab596219
Allowed for splitting the specification over several lemmas.
skalberg
parents:
14122
diff
changeset
|
97 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
98 |
val props'' = map proc_single props' |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
99 |
val frees = map snd props'' |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
100 |
val prop = myfoldr HOLogic.mk_conj (map fst props'') |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
101 |
|
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
102 |
val (vmap, prop_thawed) = Type.varify_global [] prop |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
103 |
val thawed_prop_consts = collect_consts (prop_thawed, []) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
104 |
val (altcos, overloaded) = split_list cos |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
105 |
val (names, sconsts) = split_list altcos |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
106 |
val consts = map (Syntax.read_term_global thy) sconsts |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
107 |
val _ = not (Library.exists (not o Term.is_Const) consts) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
108 |
orelse error "Specification: Non-constant found as parameter" |
14166 | 109 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
110 |
fun proc_const c = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
111 |
let |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
112 |
val (_, c') = Type.varify_global [] c |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
113 |
val (cname,ctyp) = dest_Const c' |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
114 |
in |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
115 |
(case filter (fn t => |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
116 |
let val (name, typ) = dest_Const t |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
117 |
in name = cname andalso Sign.typ_equiv thy (typ, ctyp) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
118 |
end) thawed_prop_consts of |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
119 |
[] => |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
120 |
error ("Specification: No suitable instances of constant \"" ^ |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
121 |
Syntax.string_of_term_global thy c ^ "\" found") |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
122 |
| [cf] => unvarify_global cf vmap |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
123 |
| _ => error ("Specification: Several variations of \"" ^ |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
124 |
Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)")) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
125 |
end |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
126 |
val proc_consts = map proc_const consts |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
127 |
fun mk_exist c prop = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
128 |
let |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
129 |
val T = type_of c |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
130 |
val cname = Long_Name.base_name (fst (dest_Const c)) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
131 |
val vname = if Symbol_Pos.is_identifier cname then cname else "x" |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
132 |
in HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c, prop)) end |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
133 |
val ex_prop = fold_rev mk_exist proc_consts prop |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
134 |
val cnames = map (fst o dest_Const) proc_consts |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
135 |
fun post_process (arg as (thy,thm)) = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
136 |
let |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
137 |
fun inst_all thy v thm = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
138 |
let |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59586
diff
changeset
|
139 |
val cv = Thm.global_cterm_of thy v |
59586 | 140 |
val cT = Thm.ctyp_of_cterm cv |
60801 | 141 |
val spec' = Thm.instantiate' [SOME cT] [NONE, SOME cv] spec |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
142 |
in thm RS spec' end |
60362 | 143 |
fun remove_alls frees (thm, thy) = (fold (inst_all thy) frees thm, thy) |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
144 |
fun process_single ((name, atts), rew_imp, frees) args = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
145 |
let |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
146 |
fun undo_imps thm = Thm.equal_elim (Thm.symmetric rew_imp) thm |
14164
8c3fab596219
Allowed for splitting the specification over several lemmas.
skalberg
parents:
14122
diff
changeset
|
147 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
148 |
fun add_final (thm, thy) = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
149 |
if name = "" |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
150 |
then (thm, thy) |
61268 | 151 |
else (writeln (" " ^ name ^ ": " ^ Thm.string_of_thm_global thy thm); |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
152 |
Global_Theory.store_thm (Binding.name name, thm) thy) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
153 |
in |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
154 |
swap args |
60362 | 155 |
|> remove_alls frees |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
156 |
|> apfst undo_imps |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
157 |
|> apfst Drule.export_without_context |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
158 |
|-> Thm.theory_attributes |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
159 |
(map (Attrib.attribute_cmd_global thy) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
160 |
(@{attributes [nitpick_choice_spec]} @ atts)) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
161 |
|> add_final |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
162 |
|> swap |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
163 |
end |
14164
8c3fab596219
Allowed for splitting the specification over several lemmas.
skalberg
parents:
14122
diff
changeset
|
164 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
165 |
fun process_all [proc_arg] args = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
166 |
process_single proc_arg args |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
167 |
| process_all (proc_arg::rest) (thy,thm) = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
168 |
let |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
169 |
val single_th = thm RS conjunct1 |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
170 |
val rest_th = thm RS conjunct2 |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
171 |
val (thy', _) = process_single proc_arg (thy, single_th) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
172 |
in process_all rest (thy', rest_th) end |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
173 |
| process_all [] _ = raise Fail "Choice_Specification.process_spec internal error" |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
174 |
val alt_names = map fst alt_props |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
175 |
val _ = |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
176 |
if exists (fn (name, _) => name <> "") alt_names |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
177 |
then writeln "specification" |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
178 |
else () |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
179 |
in |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
180 |
arg |
60825 | 181 |
|> apsnd (Thm.unvarify_global thy) |
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
182 |
|> process_all (zip3 alt_names rew_imps frees) |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
183 |
end |
20903 | 184 |
|
56270
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
185 |
fun after_qed [[thm]] = Proof_Context.background_theory (fn thy => |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
186 |
#1 (post_process (add_specification (zip3 names cnames overloaded) (thy, thm)))); |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
187 |
in |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
188 |
thy |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
189 |
|> Proof_Context.init_global |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
190 |
|> Variable.declare_term ex_prop |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
191 |
|> Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]] |
ce9c7a527c4b
removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents:
56254
diff
changeset
|
192 |
end |
14115 | 193 |
|
17336 | 194 |
|
14115 | 195 |
(* outer syntax *) |
196 |
||
46949 | 197 |
val opt_name = Scan.optional (Parse.name --| @{keyword ":"}) "" |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
198 |
val opt_overloaded = Parse.opt_keyword "overloaded" |
14116
63337d8e6e0f
Added optional theorem names for the constant definitions added during
skalberg
parents:
14115
diff
changeset
|
199 |
|
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46949
diff
changeset
|
200 |
val _ = |
59936
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
wenzelm
parents:
59642
diff
changeset
|
201 |
Outer_Syntax.command @{command_keyword specification} "define constants by specification" |
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46949
diff
changeset
|
202 |
(@{keyword "("} |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| @{keyword ")"} -- |
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46949
diff
changeset
|
203 |
Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop) |
56895
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
wenzelm
parents:
56270
diff
changeset
|
204 |
>> (fn (cos, alt_props) => Toplevel.theory_to_proof (process_spec cos alt_props))) |
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
205 |
|
14115 | 206 |
end |