| author | wenzelm | 
| Fri, 19 Aug 2011 18:01:23 +0200 | |
| changeset 44302 | 0a1934c5c104 | 
| parent 42425 | 2aa907d5ee4f | 
| child 44802 | 65c397cc44ec | 
| permissions | -rw-r--r-- | 
| 1526 | 1  | 
(* Title: Pure/theory.ML  | 
2  | 
Author: Lawrence C Paulson and Markus Wenzel  | 
|
3  | 
||
| 28290 | 4  | 
Logical theory content: axioms, definitions, and begin/end wrappers.  | 
| 1526 | 5  | 
*)  | 
| 16291 | 6  | 
|
| 
26668
 
65023d4fd226
removed obsolete SIGN_THEORY -- no name aliases in structure Theory;
 
wenzelm 
parents: 
26631 
diff
changeset
 | 
7  | 
signature THEORY =  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
8  | 
sig  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
9  | 
val eq_thy: theory * theory -> bool  | 
| 3996 | 10  | 
val subthy: theory * theory -> bool  | 
| 24666 | 11  | 
val assert_super: theory -> theory -> theory  | 
| 22684 | 12  | 
val parents_of: theory -> theory list  | 
13  | 
val ancestors_of: theory -> theory list  | 
|
| 42425 | 14  | 
val nodes_of: theory -> theory list  | 
| 24666 | 15  | 
val check_thy: theory -> theory_ref  | 
16  | 
val deref: theory_ref -> theory  | 
|
17  | 
val merge: theory * theory -> theory  | 
|
18  | 
val merge_refs: theory_ref * theory_ref -> theory_ref  | 
|
19  | 
val merge_list: theory list -> theory  | 
|
| 16495 | 20  | 
val checkpoint: theory -> theory  | 
21  | 
val copy: theory -> theory  | 
|
| 24666 | 22  | 
val requires: theory -> string -> string -> unit  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
23  | 
val axiom_space: theory -> Name_Space.T  | 
| 22684 | 24  | 
val axiom_table: theory -> term Symtab.table  | 
| 16339 | 25  | 
val axioms_of: theory -> (string * term) list  | 
26  | 
val all_axioms_of: theory -> (string * term) list  | 
|
| 24666 | 27  | 
val defs_of: theory -> Defs.T  | 
28  | 
val at_begin: (theory -> theory option) -> theory -> theory  | 
|
29  | 
val at_end: (theory -> theory option) -> theory -> theory  | 
|
30  | 
val begin_theory: string -> theory list -> theory  | 
|
31  | 
val end_theory: theory -> theory  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
32  | 
val add_axiom: Proof.context -> binding * term -> theory -> theory  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
33  | 
val add_deps: Proof.context -> string -> string * typ -> (string * typ) list -> theory -> theory  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
34  | 
val add_deps_global: string -> string * typ -> (string * typ) list -> theory -> theory  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
35  | 
val add_def: Proof.context -> bool -> bool -> binding * term -> theory -> theory  | 
| 29581 | 36  | 
val add_finals_i: bool -> term list -> theory -> theory  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
37  | 
val add_finals: bool -> string list -> theory -> theory  | 
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33168 
diff
changeset
 | 
38  | 
val specify_const: (binding * typ) * mixfix -> theory -> term * theory  | 
| 16495 | 39  | 
end  | 
| 1526 | 40  | 
|
| 24666 | 41  | 
structure Theory: THEORY =  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
42  | 
struct  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
43  | 
|
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
44  | 
|
| 24666 | 45  | 
(** theory context operations **)  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
46  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
47  | 
val eq_thy = Context.eq_thy;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
48  | 
val subthy = Context.subthy;  | 
| 1526 | 49  | 
|
| 
24626
 
85eceef2edc7
introduced generic concepts for theory interpretators
 
haftmann 
parents: 
24199 
diff
changeset
 | 
50  | 
fun assert_super thy1 thy2 =  | 
| 
 
85eceef2edc7
introduced generic concepts for theory interpretators
 
haftmann 
parents: 
24199 
diff
changeset
 | 
51  | 
if subthy (thy1, thy2) then thy2  | 
| 
 
85eceef2edc7
introduced generic concepts for theory interpretators
 
haftmann 
parents: 
24199 
diff
changeset
 | 
52  | 
  else raise THEORY ("Not a super theory", [thy1, thy2]);
 | 
| 
 
85eceef2edc7
introduced generic concepts for theory interpretators
 
haftmann 
parents: 
24199 
diff
changeset
 | 
53  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
54  | 
val parents_of = Context.parents_of;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
55  | 
val ancestors_of = Context.ancestors_of;  | 
| 42425 | 56  | 
fun nodes_of thy = thy :: ancestors_of thy;  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
57  | 
|
| 
24137
 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
58  | 
val check_thy = Context.check_thy;  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
59  | 
val deref = Context.deref;  | 
| 24666 | 60  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
61  | 
val merge = Context.merge;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
62  | 
val merge_refs = Context.merge_refs;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
63  | 
|
| 23600 | 64  | 
fun merge_list [] = raise THEORY ("Empty merge of theories", [])
 | 
| 21608 | 65  | 
| merge_list (thy :: thys) = Library.foldl merge (thy, thys);  | 
66  | 
||
| 16495 | 67  | 
val checkpoint = Context.checkpoint_thy;  | 
68  | 
val copy = Context.copy_thy;  | 
|
69  | 
||
| 24666 | 70  | 
fun requires thy name what =  | 
| 42425 | 71  | 
if exists (fn thy' => Context.theory_name thy' = name) (nodes_of thy) then ()  | 
| 24666 | 72  | 
  else error ("Require theory " ^ quote name ^ " as an ancestor for " ^ what);
 | 
73  | 
||
74  | 
||
75  | 
||
| 25059 | 76  | 
(** datatype thy **)  | 
| 24666 | 77  | 
|
78  | 
type wrapper = (theory -> theory option) * stamp;  | 
|
79  | 
||
80  | 
fun apply_wrappers (wrappers: wrapper list) =  | 
|
| 25059 | 81  | 
perhaps (perhaps_loop (perhaps_apply (map fst wrappers)));  | 
| 24666 | 82  | 
|
83  | 
datatype thy = Thy of  | 
|
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
84  | 
 {axioms: term Name_Space.table,
 | 
| 24666 | 85  | 
defs: Defs.T,  | 
86  | 
wrappers: wrapper list * wrapper list};  | 
|
87  | 
||
| 28290 | 88  | 
fun make_thy (axioms, defs, wrappers) = Thy {axioms = axioms, defs = defs, wrappers = wrappers};
 | 
| 24666 | 89  | 
|
| 42016 | 90  | 
structure Thy = Theory_Data_PP  | 
| 24666 | 91  | 
(  | 
92  | 
type T = thy;  | 
|
| 33159 | 93  | 
val empty_axioms = Name_Space.empty_table "axiom" : term Name_Space.table;  | 
| 33096 | 94  | 
val empty = make_thy (empty_axioms, Defs.empty, ([], []));  | 
| 24666 | 95  | 
|
| 33096 | 96  | 
  fun extend (Thy {axioms = _, defs, wrappers}) = make_thy (empty_axioms, defs, wrappers);
 | 
| 24666 | 97  | 
|
98  | 
fun merge pp (thy1, thy2) =  | 
|
99  | 
let  | 
|
| 42389 | 100  | 
val ctxt = Syntax.init_pretty pp;  | 
| 28290 | 101  | 
      val Thy {axioms = _, defs = defs1, wrappers = (bgs1, ens1)} = thy1;
 | 
102  | 
      val Thy {axioms = _, defs = defs2, wrappers = (bgs2, ens2)} = thy2;
 | 
|
| 24666 | 103  | 
|
| 33096 | 104  | 
val axioms' = empty_axioms;  | 
| 42389 | 105  | 
val defs' = Defs.merge ctxt (defs1, defs2);  | 
| 24666 | 106  | 
val bgs' = Library.merge (eq_snd op =) (bgs1, bgs2);  | 
107  | 
val ens' = Library.merge (eq_snd op =) (ens1, ens2);  | 
|
| 28290 | 108  | 
in make_thy (axioms', defs', (bgs', ens')) end;  | 
| 24666 | 109  | 
);  | 
110  | 
||
| 42016 | 111  | 
fun rep_theory thy = Thy.get thy |> (fn Thy args => args);  | 
| 24666 | 112  | 
|
| 42016 | 113  | 
fun map_thy f = Thy.map (fn (Thy {axioms, defs, wrappers}) =>
 | 
| 28290 | 114  | 
make_thy (f (axioms, defs, wrappers)));  | 
| 24666 | 115  | 
|
116  | 
||
| 28290 | 117  | 
fun map_axioms f = map_thy (fn (axioms, defs, wrappers) => (f axioms, defs, wrappers));  | 
118  | 
fun map_defs f = map_thy (fn (axioms, defs, wrappers) => (axioms, f defs, wrappers));  | 
|
119  | 
fun map_wrappers f = map_thy (fn (axioms, defs, wrappers) => (axioms, defs, f wrappers));  | 
|
| 24666 | 120  | 
|
121  | 
||
122  | 
(* basic operations *)  | 
|
123  | 
||
124  | 
val axiom_space = #1 o #axioms o rep_theory;  | 
|
125  | 
val axiom_table = #2 o #axioms o rep_theory;  | 
|
126  | 
||
127  | 
val axioms_of = Symtab.dest o #2 o #axioms o rep_theory;  | 
|
| 42425 | 128  | 
fun all_axioms_of thy = maps axioms_of (nodes_of thy);  | 
| 24666 | 129  | 
|
130  | 
val defs_of = #defs o rep_theory;  | 
|
131  | 
||
132  | 
||
133  | 
(* begin/end theory *)  | 
|
134  | 
||
135  | 
val begin_wrappers = rev o #1 o #wrappers o rep_theory;  | 
|
136  | 
val end_wrappers = rev o #2 o #wrappers o rep_theory;  | 
|
137  | 
||
138  | 
fun at_begin f = map_wrappers (apfst (cons (f, stamp ())));  | 
|
139  | 
fun at_end f = map_wrappers (apsnd (cons (f, stamp ())));  | 
|
140  | 
||
141  | 
fun begin_theory name imports =  | 
|
142  | 
let  | 
|
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
143  | 
val thy = Context.begin_thy Context.pretty_global name imports;  | 
| 24666 | 144  | 
val wrappers = begin_wrappers thy;  | 
| 33168 | 145  | 
in  | 
146  | 
thy  | 
|
147  | 
|> Sign.local_path  | 
|
148  | 
|> Sign.map_naming (Name_Space.set_theory_name name)  | 
|
149  | 
|> apply_wrappers wrappers  | 
|
150  | 
end;  | 
|
| 24666 | 151  | 
|
152  | 
fun end_theory thy =  | 
|
153  | 
thy |> apply_wrappers (end_wrappers thy) |> Context.finish_thy;  | 
|
154  | 
||
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
155  | 
|
| 3996 | 156  | 
|
| 
35985
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
157  | 
(** primitive specifications **)  | 
| 3814 | 158  | 
|
| 
35985
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
159  | 
(* raw axioms *)  | 
| 1526 | 160  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
161  | 
fun cert_axm ctxt (b, raw_tm) =  | 
| 1526 | 162  | 
let  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
163  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
32789
 
d89327de0b3c
removed redundant Sign.certify_prop, use Sign.cert_prop instead;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
164  | 
val t = Sign.cert_prop thy raw_tm  | 
| 2979 | 165  | 
handle TYPE (msg, _, _) => error msg  | 
| 16291 | 166  | 
| TERM (msg, _) => error msg;  | 
| 
35987
 
7c728daf4876
disallow sort constraints in primitive Theory.add_axiom/add_def -- handled in Thm.add_axiom/add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
167  | 
val _ = Term.no_dummy_patterns t handle TERM (msg, _) => error msg;  | 
| 
 
7c728daf4876
disallow sort constraints in primitive Theory.add_axiom/add_def -- handled in Thm.add_axiom/add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
168  | 
|
| 
 
7c728daf4876
disallow sort constraints in primitive Theory.add_axiom/add_def -- handled in Thm.add_axiom/add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
169  | 
val bad_sorts =  | 
| 
 
7c728daf4876
disallow sort constraints in primitive Theory.add_axiom/add_def -- handled in Thm.add_axiom/add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
170  | 
rev ((fold_types o fold_atyps_sorts) (fn (_, []) => I | (T, _) => insert (op =) T) t []);  | 
| 
 
7c728daf4876
disallow sort constraints in primitive Theory.add_axiom/add_def -- handled in Thm.add_axiom/add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
171  | 
val _ = null bad_sorts orelse  | 
| 
 
7c728daf4876
disallow sort constraints in primitive Theory.add_axiom/add_def -- handled in Thm.add_axiom/add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
172  | 
      error ("Illegal sort constraints in primitive specification: " ^
 | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
173  | 
commas (map (Syntax.string_of_typ (Config.put show_sorts true ctxt)) bad_sorts));  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
174  | 
in (b, Sign.no_vars ctxt t) end  | 
| 
42381
 
309ec68442c6
added Binding.print convenience, which includes quote already;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
175  | 
  handle ERROR msg => cat_error msg ("The error(s) above occurred in axiom " ^ Binding.print b);
 | 
| 1526 | 176  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
177  | 
fun add_axiom ctxt raw_axm thy = thy |> map_axioms (fn axioms =>  | 
| 1526 | 178  | 
let  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
179  | 
val axm = apsnd Logic.varify_global (cert_axm ctxt raw_axm);  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
180  | 
val (_, axioms') = Name_Space.define ctxt true (Sign.naming_of thy) axm axioms;  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
181  | 
in axioms' end);  | 
| 1526 | 182  | 
|
183  | 
||
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
184  | 
(* dependencies *)  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
185  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
186  | 
fun dependencies ctxt unchecked def description lhs rhs =  | 
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
187  | 
let  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
188  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
189  | 
val consts = Sign.consts_of thy;  | 
| 19727 | 190  | 
fun prep const =  | 
| 
39133
 
70d3915c92f0
pretty printing: prefer regular Proof.context over Pretty.pp, which is mostly for special bootstrap purposes involving theory merge, for example;
 
wenzelm 
parents: 
36610 
diff
changeset
 | 
191  | 
let val Const (c, T) = Sign.no_vars ctxt (Const const)  | 
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
34259 
diff
changeset
 | 
192  | 
in (c, Consts.typargs consts (c, Logic.varifyT_global T)) end;  | 
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
193  | 
|
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
194  | 
val lhs_vars = Term.add_tfreesT (#2 lhs) [];  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
195  | 
val rhs_extras = fold (#2 #> Term.fold_atyps (fn TFree v =>  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
196  | 
if member (op =) lhs_vars v then I else insert (op =) v | _ => I)) rhs [];  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
197  | 
val _ =  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
198  | 
if null rhs_extras then ()  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
199  | 
      else error ("Specification depends on extra type variables: " ^
 | 
| 
39133
 
70d3915c92f0
pretty printing: prefer regular Proof.context over Pretty.pp, which is mostly for special bootstrap purposes involving theory merge, for example;
 
wenzelm 
parents: 
36610 
diff
changeset
 | 
200  | 
commas_quote (map (Syntax.string_of_typ ctxt o TFree) rhs_extras) ^  | 
| 
33701
 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 
wenzelm 
parents: 
33173 
diff
changeset
 | 
201  | 
"\nThe error(s) above occurred in " ^ quote description);  | 
| 42384 | 202  | 
in Defs.define ctxt unchecked def description (prep lhs) (map prep rhs) end;  | 
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
203  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
204  | 
fun add_deps ctxt a raw_lhs raw_rhs thy =  | 
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
205  | 
let  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
206  | 
val lhs :: rhs = map (dest_Const o Sign.cert_term thy o Const) (raw_lhs :: raw_rhs);  | 
| 
33701
 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 
wenzelm 
parents: 
33173 
diff
changeset
 | 
207  | 
val description = if a = "" then #1 lhs ^ " axiom" else a;  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
208  | 
in thy |> map_defs (dependencies ctxt false NONE description lhs rhs) end;  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
209  | 
|
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
210  | 
fun add_deps_global a x y thy = add_deps (Syntax.init_pretty_global thy) a x y thy;  | 
| 17706 | 211  | 
|
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33168 
diff
changeset
 | 
212  | 
fun specify_const decl thy =  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
213  | 
let val (t as Const const, thy') = Sign.declare_const_global decl thy;  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
214  | 
in (t, add_deps_global "" const [] thy') end;  | 
| 25017 | 215  | 
|
| 17706 | 216  | 
|
| 
35985
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
217  | 
(* overloading *)  | 
| 9280 | 218  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
219  | 
fun check_overloading ctxt overloaded (c, T) =  | 
| 16291 | 220  | 
let  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
221  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
222  | 
|
| 24763 | 223  | 
val declT = Sign.the_const_constraint thy c  | 
224  | 
handle TYPE (msg, _, _) => error msg;  | 
|
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
34259 
diff
changeset
 | 
225  | 
val T' = Logic.varifyT_global T;  | 
| 16944 | 226  | 
|
| 
39134
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
39133 
diff
changeset
 | 
227  | 
fun message sorts txt =  | 
| 16944 | 228  | 
[Pretty.block [Pretty.str "Specification of constant ",  | 
| 
39134
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
39133 
diff
changeset
 | 
229  | 
Pretty.str c, Pretty.str " ::", Pretty.brk 1,  | 
| 
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
39133 
diff
changeset
 | 
230  | 
Pretty.quote (Syntax.pretty_typ (Config.put show_sorts sorts ctxt) T)],  | 
| 16944 | 231  | 
Pretty.str txt] |> Pretty.chunks |> Pretty.string_of;  | 
| 16291 | 232  | 
in  | 
| 16944 | 233  | 
if Sign.typ_instance thy (declT, T') then ()  | 
234  | 
else if Type.raw_instance (declT, T') then  | 
|
| 
39134
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
39133 
diff
changeset
 | 
235  | 
error (message true "imposes additional sort constraints on the constant declaration")  | 
| 16944 | 236  | 
else if overloaded then ()  | 
| 
39134
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
39133 
diff
changeset
 | 
237  | 
else warning (message false "is strictly less general than the declared type")  | 
| 9280 | 238  | 
end;  | 
239  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
240  | 
|
| 
35985
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
241  | 
(* definitional axioms *)  | 
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
242  | 
|
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
243  | 
local  | 
| 16291 | 244  | 
|
| 
42394
 
c65c07d9967a
recovered Theory.check_def: full name needs to be determined from background thy, not auxiliary ctxt (broken in 774df7c59508, caused Nitpick.all_axioms_of to produce bad results);
 
wenzelm 
parents: 
42389 
diff
changeset
 | 
245  | 
fun check_def ctxt thy unchecked overloaded (b, tm) defs =  | 
| 16291 | 246  | 
let  | 
| 29581 | 247  | 
val name = Sign.full_name thy b;  | 
| 
35988
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35987 
diff
changeset
 | 
248  | 
val ((lhs, rhs), _) = Primitive_Defs.dest_def ctxt Term.is_Const (K false) (K false) tm  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35987 
diff
changeset
 | 
249  | 
handle TERM (msg, _) => error msg;  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35987 
diff
changeset
 | 
250  | 
val lhs_const = Term.dest_Const (Term.head_of lhs);  | 
| 16944 | 251  | 
val rhs_consts = fold_aterms (fn Const const => insert (op =) const | _ => I) rhs [];  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
252  | 
val _ = check_overloading ctxt overloaded lhs_const;  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
253  | 
in defs |> dependencies ctxt unchecked (SOME name) name lhs_const rhs_consts end  | 
| 18678 | 254  | 
handle ERROR msg => cat_error msg (Pretty.string_of (Pretty.block  | 
| 
42381
 
309ec68442c6
added Binding.print convenience, which includes quote already;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
255  | 
   [Pretty.str ("The error(s) above occurred in definition " ^ Binding.print b ^ ":"),
 | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
256  | 
Pretty.fbrk, Pretty.quote (Syntax.pretty_term ctxt tm)]));  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
257  | 
|
| 16291 | 258  | 
in  | 
259  | 
||
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
260  | 
fun add_def ctxt unchecked overloaded raw_axm thy =  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
261  | 
let val axm = cert_axm ctxt raw_axm in  | 
| 
35985
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
262  | 
thy  | 
| 
42394
 
c65c07d9967a
recovered Theory.check_def: full name needs to be determined from background thy, not auxiliary ctxt (broken in 774df7c59508, caused Nitpick.all_axioms_of to produce bad results);
 
wenzelm 
parents: 
42389 
diff
changeset
 | 
263  | 
|> map_defs (check_def ctxt thy unchecked overloaded axm)  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
264  | 
|> add_axiom ctxt axm  | 
| 
35985
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
265  | 
end;  | 
| 16291 | 266  | 
|
267  | 
end;  | 
|
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
268  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
269  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
270  | 
(* add_finals(_i) *)  | 
| 
14223
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
14204 
diff
changeset
 | 
271  | 
|
| 16291 | 272  | 
local  | 
273  | 
||
| 17706 | 274  | 
fun gen_add_finals prep_term overloaded args thy =  | 
| 
14223
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
14204 
diff
changeset
 | 
275  | 
let  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
276  | 
val ctxt = Syntax.init_pretty_global thy;  | 
| 17706 | 277  | 
fun const_of (Const const) = const  | 
278  | 
| const_of (Free _) = error "Attempt to finalize variable (or undeclared constant)"  | 
|
279  | 
| const_of _ = error "Attempt to finalize non-constant term";  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
280  | 
fun specify (c, T) = dependencies ctxt false NONE (c ^ " axiom") (c, T) [];  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
281  | 
val finalize =  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
282  | 
specify o tap (check_overloading ctxt overloaded) o const_of o  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
283  | 
Sign.cert_term thy o prep_term ctxt;  | 
| 17706 | 284  | 
in thy |> map_defs (fold finalize args) end;  | 
| 16291 | 285  | 
|
| 
14223
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
14204 
diff
changeset
 | 
286  | 
in  | 
| 16291 | 287  | 
|
| 29581 | 288  | 
val add_finals_i = gen_add_finals (K I);  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
289  | 
val add_finals = gen_add_finals Syntax.read_term;  | 
| 16291 | 290  | 
|
| 
14223
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
14204 
diff
changeset
 | 
291  | 
end;  | 
| 
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
14204 
diff
changeset
 | 
292  | 
|
| 1526 | 293  | 
end;  |