| author | haftmann | 
| Tue, 31 Oct 2006 14:58:12 +0100 | |
| changeset 21125 | 9b7d35ca1eef | 
| parent 20549 | c643984eb94b | 
| child 21608 | 2ca27eeb2841 | 
| permissions | -rw-r--r-- | 
| 1526 | 1  | 
(* Title: Pure/theory.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson and Markus Wenzel  | 
|
4  | 
||
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
5  | 
Logical theory content: axioms, definitions, oracles.  | 
| 1526 | 6  | 
*)  | 
| 16291 | 7  | 
|
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
8  | 
signature BASIC_THEORY =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
9  | 
sig  | 
| 16495 | 10  | 
val sign_of: theory -> theory (*obsolete*)  | 
| 3996 | 11  | 
val rep_theory: theory ->  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
12  | 
   {axioms: term NameSpace.table,
 | 
| 17706 | 13  | 
defs: Defs.T,  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
14  | 
oracles: ((theory * Object.T -> term) * stamp) NameSpace.table}  | 
| 3996 | 15  | 
val parents_of: theory -> theory list  | 
| 4019 | 16  | 
val ancestors_of: theory -> theory list  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
17  | 
val eq_thy: theory * theory -> bool  | 
| 3996 | 18  | 
val subthy: theory * theory -> bool  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
19  | 
val cert_axm: theory -> string * term -> string * term  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
20  | 
val read_def_axm: theory * (indexname -> typ option) * (indexname -> sort option) ->  | 
| 6311 | 21  | 
string list -> string * string -> string * term  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
22  | 
val read_axm: theory -> string * string -> string * term  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
23  | 
val inferT_axm: theory -> string * term -> string * term  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
24  | 
end  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
25  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
26  | 
signature THEORY =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
27  | 
sig  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
28  | 
include BASIC_THEORY  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
29  | 
include SIGN_THEORY  | 
| 16495 | 30  | 
val begin_theory: string -> theory list -> theory  | 
31  | 
val end_theory: theory -> theory  | 
|
32  | 
val checkpoint: theory -> theory  | 
|
33  | 
val copy: theory -> theory  | 
|
| 16536 | 34  | 
val init_data: theory -> theory  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
35  | 
val axiom_space: theory -> NameSpace.T  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
36  | 
val oracle_space: theory -> NameSpace.T  | 
| 16339 | 37  | 
val axioms_of: theory -> (string * term) list  | 
38  | 
val all_axioms_of: theory -> (string * term) list  | 
|
| 17706 | 39  | 
val defs_of : theory -> Defs.T  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
40  | 
val self_ref: theory -> theory_ref  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
41  | 
val deref: theory_ref -> theory  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
42  | 
val merge: theory * theory -> theory (*exception TERM*)  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
43  | 
val merge_refs: theory_ref * theory_ref -> theory_ref (*exception TERM*)  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
44  | 
val requires: theory -> string -> string -> unit  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
45  | 
val assert_super: theory -> theory -> theory  | 
| 3996 | 46  | 
val add_axioms: (bstring * string) list -> theory -> theory  | 
47  | 
val add_axioms_i: (bstring * term) list -> theory -> theory  | 
|
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
48  | 
val add_deps: string -> string * typ -> (string * typ) list -> theory -> theory  | 
| 19630 | 49  | 
val add_defs: bool -> bool -> (bstring * string) list -> theory -> theory  | 
50  | 
val add_defs_i: bool -> bool -> (bstring * term) list -> theory -> theory  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
51  | 
val add_finals: bool -> string list -> theory -> theory  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
52  | 
val add_finals_i: bool -> term list -> theory -> theory  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
53  | 
val add_oracle: bstring * (theory * Object.T -> term) -> theory -> theory  | 
| 16495 | 54  | 
end  | 
| 1526 | 55  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
56  | 
structure Theory: THEORY =  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
57  | 
struct  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
58  | 
|
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
59  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
60  | 
(** type theory **)  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
61  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
62  | 
(* context operations *)  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
63  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
64  | 
val eq_thy = Context.eq_thy;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
65  | 
val subthy = Context.subthy;  | 
| 1526 | 66  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
67  | 
val parents_of = Context.parents_of;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
68  | 
val ancestors_of = Context.ancestors_of;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
69  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
70  | 
val self_ref = Context.self_ref;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
71  | 
val deref = Context.deref;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
72  | 
val merge = Context.merge;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
73  | 
val merge_refs = Context.merge_refs;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
74  | 
|
| 16495 | 75  | 
val begin_theory = Sign.local_path oo Context.begin_thy Sign.pp;  | 
76  | 
val end_theory = Context.finish_thy;  | 
|
77  | 
val checkpoint = Context.checkpoint_thy;  | 
|
78  | 
val copy = Context.copy_thy;  | 
|
79  | 
||
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
80  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
81  | 
(* signature operations *)  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
82  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
83  | 
val sign_of = I;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
84  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
85  | 
structure SignTheory: SIGN_THEORY = Sign;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
86  | 
open SignTheory;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
87  | 
|
| 2206 | 88  | 
|
| 3996 | 89  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
90  | 
(** datatype thy **)  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
91  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
92  | 
datatype thy = Thy of  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
93  | 
 {axioms: term NameSpace.table,
 | 
| 17706 | 94  | 
defs: Defs.T,  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
95  | 
oracles: ((theory * Object.T -> term) * stamp) NameSpace.table};  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
96  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
97  | 
fun make_thy (axioms, defs, oracles) =  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
98  | 
  Thy {axioms = axioms, defs = defs, oracles = oracles};
 | 
| 1526 | 99  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
100  | 
fun err_dup_axms dups = error ("Duplicate axiom(s): " ^ commas_quote dups);
 | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
101  | 
fun err_dup_oras dups = error ("Duplicate oracle(s): " ^ commas_quote dups);
 | 
| 3996 | 102  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
103  | 
structure ThyData = TheoryDataFun  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
104  | 
(struct  | 
| 16536 | 105  | 
val name = "Pure/theory";  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
106  | 
type T = thy;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
107  | 
val empty = make_thy (NameSpace.empty_table, Defs.empty, NameSpace.empty_table);  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
108  | 
val copy = I;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
109  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
110  | 
  fun extend (Thy {axioms, defs, oracles}) = make_thy (NameSpace.empty_table, defs, oracles);
 | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
111  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
112  | 
fun merge pp (thy1, thy2) =  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
113  | 
let  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
114  | 
      val Thy {axioms = _, defs = defs1, oracles = oracles1} = thy1;
 | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
115  | 
      val Thy {axioms = _, defs = defs2, oracles = oracles2} = thy2;
 | 
| 1526 | 116  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
117  | 
val axioms = NameSpace.empty_table;  | 
| 19693 | 118  | 
val defs = Defs.merge pp (defs1, defs2);  | 
| 17496 | 119  | 
val oracles = NameSpace.merge_tables (eq_snd (op =)) (oracles1, oracles2)  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
120  | 
handle Symtab.DUPS dups => err_dup_oras dups;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
121  | 
in make_thy (axioms, defs, oracles) end;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
122  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
123  | 
fun print _ _ = ();  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
124  | 
end);  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
125  | 
|
| 16536 | 126  | 
val init_data = ThyData.init;  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
127  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
128  | 
fun rep_theory thy = ThyData.get thy |> (fn Thy args => args);  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
129  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
130  | 
fun map_thy f = ThyData.map (fn (Thy {axioms, defs, oracles}) =>
 | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
131  | 
make_thy (f (axioms, defs, oracles)));  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
132  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
133  | 
fun map_axioms f = map_thy (fn (axioms, defs, oracles) => (f axioms, defs, oracles));  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
134  | 
fun map_defs f = map_thy (fn (axioms, defs, oracles) => (axioms, f defs, oracles));  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
135  | 
fun map_oracles f = map_thy (fn (axioms, defs, oracles) => (axioms, defs, f oracles));  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
136  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
137  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
138  | 
(* basic operations *)  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
139  | 
|
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
140  | 
val axiom_space = #1 o #axioms o rep_theory;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
141  | 
val oracle_space = #1 o #oracles o rep_theory;  | 
| 3996 | 142  | 
|
| 16339 | 143  | 
val axioms_of = Symtab.dest o #2 o #axioms o rep_theory;  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19428 
diff
changeset
 | 
144  | 
fun all_axioms_of thy = maps axioms_of (thy :: ancestors_of thy);  | 
| 16339 | 145  | 
|
| 16803 | 146  | 
val defs_of = #defs o rep_theory;  | 
| 
16743
 
21dbff595bf6
1) all theorems in Orderings can now be given as a parameter
 
obua 
parents: 
16600 
diff
changeset
 | 
147  | 
|
| 4970 | 148  | 
fun requires thy name what =  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
149  | 
if Context.exists_name name thy then ()  | 
| 4846 | 150  | 
  else error ("Require theory " ^ quote name ^ " as an ancestor for " ^ what);
 | 
| 1526 | 151  | 
|
| 6369 | 152  | 
fun assert_super thy1 thy2 =  | 
153  | 
if subthy (thy1, thy2) then thy2  | 
|
154  | 
  else raise THEORY ("Not a super theory", [thy1, thy2]);
 | 
|
155  | 
||
| 3996 | 156  | 
|
| 6311 | 157  | 
|
| 3814 | 158  | 
(** add axioms **)  | 
159  | 
||
| 1526 | 160  | 
(* prepare axioms *)  | 
161  | 
||
| 18678 | 162  | 
fun err_in_axm msg name =  | 
163  | 
  cat_error msg ("The error(s) above occurred in axiom " ^ quote name);
 | 
|
| 1526 | 164  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
165  | 
fun cert_axm thy (name, raw_tm) =  | 
| 1526 | 166  | 
let  | 
| 
18968
 
52639ad19a96
adapted Sign.infer_types(_simult), Sign.certify_term/prop;
 
wenzelm 
parents: 
18943 
diff
changeset
 | 
167  | 
val (t, T, _) = Sign.certify_prop thy raw_tm  | 
| 2979 | 168  | 
handle TYPE (msg, _, _) => error msg  | 
| 16291 | 169  | 
| TERM (msg, _) => error msg;  | 
| 1526 | 170  | 
in  | 
| 9537 | 171  | 
Term.no_dummy_patterns t handle TERM (msg, _) => error msg;  | 
| 
18968
 
52639ad19a96
adapted Sign.infer_types(_simult), Sign.certify_term/prop;
 
wenzelm 
parents: 
18943 
diff
changeset
 | 
172  | 
(name, Sign.no_vars (Sign.pp thy) t)  | 
| 9629 | 173  | 
end;  | 
| 1526 | 174  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
175  | 
fun read_def_axm (thy, types, sorts) used (name, str) =  | 
| 3814 | 176  | 
let  | 
| 18857 | 177  | 
val ts = Syntax.read (Context.Theory thy) (Sign.is_logtype thy) (Sign.syn_of thy) propT str;  | 
| 
18968
 
52639ad19a96
adapted Sign.infer_types(_simult), Sign.certify_term/prop;
 
wenzelm 
parents: 
18943 
diff
changeset
 | 
178  | 
val (t, _) =  | 
| 20155 | 179  | 
Sign.infer_types (Sign.pp thy) thy (Sign.consts_of thy)  | 
180  | 
types sorts (Name.make_context used) true (ts, propT);  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
181  | 
in cert_axm thy (name, t) end  | 
| 18678 | 182  | 
handle ERROR msg => err_in_axm msg name;  | 
| 1526 | 183  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
184  | 
fun read_axm thy name_str = read_def_axm (thy, K NONE, K NONE) [] name_str;  | 
| 5057 | 185  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
186  | 
fun inferT_axm thy (name, pre_tm) =  | 
| 16291 | 187  | 
let  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
188  | 
val pp = Sign.pp thy;  | 
| 
18968
 
52639ad19a96
adapted Sign.infer_types(_simult), Sign.certify_term/prop;
 
wenzelm 
parents: 
18943 
diff
changeset
 | 
189  | 
val (t, _) =  | 
| 20155 | 190  | 
Sign.infer_types pp thy (Sign.consts_of thy)  | 
191  | 
(K NONE) (K NONE) Name.context true ([pre_tm], propT);  | 
|
| 18943 | 192  | 
in (name, Sign.no_vars pp t) end  | 
| 18678 | 193  | 
handle ERROR msg => err_in_axm msg name;  | 
| 1526 | 194  | 
|
195  | 
||
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
196  | 
(* add_axioms(_i) *)  | 
| 1526 | 197  | 
|
| 16291 | 198  | 
local  | 
199  | 
||
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
200  | 
fun gen_add_axioms prep_axm raw_axms thy = thy |> map_axioms (fn axioms =>  | 
| 1526 | 201  | 
let  | 
| 16991 | 202  | 
val axms = map (apsnd (Compress.term thy o Logic.varify) o prep_axm thy) raw_axms;  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
203  | 
val axioms' = NameSpace.extend_table (Sign.naming_of thy) (axioms, axms)  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
204  | 
handle Symtab.DUPS dups => err_dup_axms dups;  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
205  | 
in axioms' end);  | 
| 1526 | 206  | 
|
| 16291 | 207  | 
in  | 
208  | 
||
209  | 
val add_axioms = gen_add_axioms read_axm;  | 
|
210  | 
val add_axioms_i = gen_add_axioms cert_axm;  | 
|
211  | 
||
212  | 
end;  | 
|
| 1526 | 213  | 
|
214  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
215  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
216  | 
(** add constant definitions **)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
217  | 
|
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
218  | 
(* dependencies *)  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
219  | 
|
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
220  | 
fun dependencies thy unchecked is_def name lhs rhs =  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
221  | 
let  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
222  | 
val pp = Sign.pp thy;  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
223  | 
val consts = Sign.consts_of thy;  | 
| 19727 | 224  | 
fun prep const =  | 
225  | 
let val Const (c, T) = Sign.no_vars pp (Const const)  | 
|
| 19806 | 226  | 
in (c, Consts.typargs consts (c, Compress.typ thy (Logic.varifyT T))) end;  | 
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
227  | 
|
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
228  | 
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
 | 
229  | 
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
 | 
230  | 
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
 | 
231  | 
val _ =  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
232  | 
if null rhs_extras then ()  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
233  | 
      else error ("Specification depends on extra type variables: " ^
 | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
234  | 
commas_quote (map (Pretty.string_of_typ pp o TFree) rhs_extras) ^  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
235  | 
"\nThe error(s) above occurred in " ^ quote name);  | 
| 19727 | 236  | 
in Defs.define pp unchecked is_def (Context.theory_name thy) name (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
 | 
237  | 
|
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
238  | 
fun add_deps a raw_lhs raw_rhs thy =  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
239  | 
let  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
240  | 
val lhs :: rhs = map (dest_Const o Sign.cert_term thy o Const) (raw_lhs :: raw_rhs);  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
241  | 
val name = if a = "" then (#1 lhs ^ " axiom") else a;  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
242  | 
in thy |> map_defs (dependencies thy false false name lhs rhs) end;  | 
| 17706 | 243  | 
|
244  | 
||
| 16944 | 245  | 
(* check_overloading *)  | 
| 9280 | 246  | 
|
| 16944 | 247  | 
fun check_overloading thy overloaded (c, T) =  | 
| 16291 | 248  | 
let  | 
| 16944 | 249  | 
val declT =  | 
250  | 
(case Sign.const_constraint thy c of  | 
|
251  | 
        NONE => error ("Undeclared constant " ^ quote c)
 | 
|
252  | 
| SOME declT => declT);  | 
|
| 19806 | 253  | 
val T' = Logic.varifyT T;  | 
| 16944 | 254  | 
|
255  | 
fun message txt =  | 
|
256  | 
[Pretty.block [Pretty.str "Specification of constant ",  | 
|
257  | 
Pretty.str c, Pretty.str " ::", Pretty.brk 1, Pretty.quote (Sign.pretty_typ thy T)],  | 
|
258  | 
Pretty.str txt] |> Pretty.chunks |> Pretty.string_of;  | 
|
| 16291 | 259  | 
in  | 
| 16944 | 260  | 
if Sign.typ_instance thy (declT, T') then ()  | 
261  | 
else if Type.raw_instance (declT, T') then  | 
|
262  | 
error (Library.setmp show_sorts true  | 
|
263  | 
message "imposes additional sort constraints on the constant declaration")  | 
|
264  | 
else if overloaded then ()  | 
|
265  | 
else warning (message "is strictly less general than the declared type");  | 
|
266  | 
(c, T)  | 
|
| 9280 | 267  | 
end;  | 
268  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
269  | 
|
| 16291 | 270  | 
(* check_def *)  | 
271  | 
||
| 19630 | 272  | 
fun check_def thy unchecked overloaded (bname, tm) defs =  | 
| 16291 | 273  | 
let  | 
| 17706 | 274  | 
val name = Sign.full_name thy bname;  | 
| 19693 | 275  | 
val (lhs_const, rhs) = Sign.cert_def (Sign.pp thy) tm;  | 
| 16944 | 276  | 
val rhs_consts = fold_aterms (fn Const const => insert (op =) const | _ => I) rhs [];  | 
| 18943 | 277  | 
val _ = check_overloading thy overloaded lhs_const;  | 
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
278  | 
in defs |> dependencies thy unchecked true name lhs_const rhs_consts end  | 
| 18678 | 279  | 
handle ERROR msg => cat_error msg (Pretty.string_of (Pretty.block  | 
| 16883 | 280  | 
   [Pretty.str ("The error(s) above occurred in definition " ^ quote bname ^ ":"),
 | 
| 19693 | 281  | 
Pretty.fbrk, Pretty.quote (Sign.pretty_term thy tm)]));  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
282  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
283  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
284  | 
(* add_defs(_i) *)  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
285  | 
|
| 16291 | 286  | 
local  | 
| 9320 | 287  | 
|
| 19630 | 288  | 
fun gen_add_defs prep_axm unchecked overloaded raw_axms thy =  | 
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
289  | 
let val axms = map (prep_axm thy) raw_axms in  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
290  | 
thy  | 
| 19630 | 291  | 
|> map_defs (fold (check_def thy unchecked overloaded) axms)  | 
| 9320 | 292  | 
|> add_axioms_i axms  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
293  | 
end;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
294  | 
|
| 16291 | 295  | 
in  | 
296  | 
||
297  | 
val add_defs_i = gen_add_defs cert_axm;  | 
|
298  | 
val add_defs = gen_add_defs read_axm;  | 
|
299  | 
||
300  | 
end;  | 
|
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
301  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
302  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
303  | 
(* add_finals(_i) *)  | 
| 
14223
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
14204 
diff
changeset
 | 
304  | 
|
| 16291 | 305  | 
local  | 
306  | 
||
| 17706 | 307  | 
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
 | 
308  | 
let  | 
| 17706 | 309  | 
fun const_of (Const const) = const  | 
310  | 
| const_of (Free _) = error "Attempt to finalize variable (or undeclared constant)"  | 
|
311  | 
| const_of _ = error "Attempt to finalize non-constant term";  | 
|
| 
19708
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
312  | 
fun specify (c, T) = dependencies thy false false (c ^ " axiom") (c, T) [];  | 
| 
 
a508bde37a81
added add_deps, which actually records dependencies of consts (unlike add_finals);
 
wenzelm 
parents: 
19700 
diff
changeset
 | 
313  | 
val finalize = specify o check_overloading thy overloaded o const_of o prep_term thy;  | 
| 17706 | 314  | 
in thy |> map_defs (fold finalize args) end;  | 
| 16291 | 315  | 
|
| 
14223
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
14204 
diff
changeset
 | 
316  | 
in  | 
| 16291 | 317  | 
|
| 17706 | 318  | 
val add_finals = gen_add_finals Sign.read_term;  | 
319  | 
val add_finals_i = gen_add_finals Sign.cert_term;  | 
|
| 16291 | 320  | 
|
| 
14223
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
14204 
diff
changeset
 | 
321  | 
end;  | 
| 
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
14204 
diff
changeset
 | 
322  | 
|
| 
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
14204 
diff
changeset
 | 
323  | 
|
| 3878 | 324  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
325  | 
(** add oracle **)  | 
| 3814 | 326  | 
|
| 
16443
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
327  | 
fun add_oracle (bname, oracle) thy = thy |> map_oracles (fn oracles =>  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
328  | 
NameSpace.extend_table (Sign.naming_of thy) (oracles, [(bname, (oracle, stamp ()))])  | 
| 
 
82a116532e3e
type theory, theory_ref, exception THEORY and related operations imported from Context;
 
wenzelm 
parents: 
16369 
diff
changeset
 | 
329  | 
handle Symtab.DUPS dups => err_dup_oras dups);  | 
| 3885 | 330  | 
|
| 1526 | 331  | 
end;  | 
332  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
333  | 
structure BasicTheory: BASIC_THEORY = Theory;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
334  | 
open BasicTheory;  |