| author | paulson | 
| Wed, 08 Jan 1997 15:17:25 +0100 | |
| changeset 2495 | 82ec47e0a8d3 | 
| parent 2385 | 73d1435aa729 | 
| child 2693 | 8300bba275e3 | 
| permissions | -rw-r--r-- | 
| 1526 | 1 | (* Title: Pure/theory.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson and Markus Wenzel | |
| 4 | Copyright 1996 University of Cambridge | |
| 5 | ||
| 6 | Theories | |
| 7 | *) | |
| 8 | ||
| 9 | signature THEORY = | |
| 10 | sig | |
| 11 | type theory | |
| 12 | exception THEORY of string * theory list | |
| 13 | val rep_theory : theory -> | |
| 1539 | 14 |     {sign: Sign.sg, oraopt: (Sign.sg * exn -> term) option,
 | 
| 15 | new_axioms: term Symtab.table, parents: theory list} | |
| 1526 | 16 | val sign_of : theory -> Sign.sg | 
| 17 | val syn_of : theory -> Syntax.syntax | |
| 18 | val stamps_of_thy : theory -> string ref list | |
| 19 | val parents_of : theory -> theory list | |
| 20 | val subthy : theory * theory -> bool | |
| 21 | val eq_thy : theory * theory -> bool | |
| 22 | val proto_pure_thy : theory | |
| 23 | val pure_thy : theory | |
| 24 | val cpure_thy : theory | |
| 25 | (*theory primitives*) | |
| 1539 | 26 | val add_classes : (class * class list) list -> theory -> theory | 
| 27 | val add_classrel : (class * class) list -> theory -> theory | |
| 28 | val add_defsort : sort -> theory -> theory | |
| 29 | val add_types : (string * int * mixfix) list -> theory -> theory | |
| 30 | val add_tyabbrs : (string * string list * string * mixfix) list | |
| 1526 | 31 | -> theory -> theory | 
| 1539 | 32 | val add_tyabbrs_i : (string * string list * typ * mixfix) list | 
| 1526 | 33 | -> theory -> theory | 
| 1539 | 34 | val add_arities : (string * sort list * sort) list -> theory -> theory | 
| 35 | val add_consts : (string * string * mixfix) list -> theory -> theory | |
| 36 | val add_consts_i : (string * typ * mixfix) list -> theory -> theory | |
| 37 | val add_syntax : (string * string * mixfix) list -> theory -> theory | |
| 38 | val add_syntax_i : (string * typ * mixfix) list -> theory -> theory | |
| 2359 | 39 | val add_modesyntax : string * bool -> (string * string * mixfix) list -> theory -> theory | 
| 40 | val add_modesyntax_i : string * bool -> (string * typ * mixfix) list -> theory -> theory | |
| 1539 | 41 | val add_trfuns : | 
| 1526 | 42 | (string * (Syntax.ast list -> Syntax.ast)) list * | 
| 43 | (string * (term list -> term)) list * | |
| 44 | (string * (term list -> term)) list * | |
| 45 | (string * (Syntax.ast list -> Syntax.ast)) list -> theory -> theory | |
| 2385 | 46 | val add_trfunsT : | 
| 47 | (string * (typ -> term list -> term)) list -> theory -> theory | |
| 1526 | 48 | val add_trrules : (string * string)Syntax.trrule list -> theory -> theory | 
| 1539 | 49 | val add_trrules_i : Syntax.ast Syntax.trrule list -> theory -> theory | 
| 1526 | 50 | val cert_axm : Sign.sg -> string * term -> string * term | 
| 51 | val read_axm : Sign.sg -> string * string -> string * term | |
| 52 | val inferT_axm : Sign.sg -> string * term -> string * term | |
| 1539 | 53 | val add_axioms : (string * string) list -> theory -> theory | 
| 54 | val add_axioms_i : (string * term) list -> theory -> theory | |
| 55 | val add_thyname : string -> theory -> theory | |
| 56 | ||
| 57 | val set_oracle : (Sign.sg * exn -> term) -> theory -> theory | |
| 1526 | 58 | |
| 59 | val merge_theories : theory * theory -> theory | |
| 60 | val merge_thy_list : bool -> theory list -> theory | |
| 61 | end; | |
| 62 | ||
| 63 | ||
| 64 | structure Theory : THEORY = | |
| 65 | struct | |
| 2206 | 66 | |
| 67 | (** datatype theory **) | |
| 1526 | 68 | |
| 69 | datatype theory = | |
| 70 |   Theory of {
 | |
| 71 | sign: Sign.sg, | |
| 1539 | 72 | oraopt: (Sign.sg * exn -> term) option, | 
| 1526 | 73 | new_axioms: term Symtab.table, | 
| 74 | parents: theory list}; | |
| 75 | ||
| 76 | fun rep_theory (Theory args) = args; | |
| 77 | ||
| 78 | (*errors involving theories*) | |
| 79 | exception THEORY of string * theory list; | |
| 80 | ||
| 81 | ||
| 82 | val sign_of = #sign o rep_theory; | |
| 83 | val syn_of = #syn o Sign.rep_sg o sign_of; | |
| 84 | ||
| 85 | (*stamps associated with a theory*) | |
| 86 | val stamps_of_thy = #stamps o Sign.rep_sg o sign_of; | |
| 87 | ||
| 88 | (*return the immediate ancestors*) | |
| 89 | val parents_of = #parents o rep_theory; | |
| 90 | ||
| 91 | ||
| 92 | (*compare theories*) | |
| 93 | val subthy = Sign.subsig o pairself sign_of; | |
| 94 | val eq_thy = Sign.eq_sg o pairself sign_of; | |
| 95 | ||
| 96 | ||
| 97 | (* the Pure theories *) | |
| 98 | ||
| 99 | val proto_pure_thy = | |
| 1539 | 100 |   Theory {sign = Sign.proto_pure, oraopt = None, 
 | 
| 101 | new_axioms = Symtab.null, parents = []}; | |
| 1526 | 102 | |
| 103 | val pure_thy = | |
| 1539 | 104 |   Theory {sign = Sign.pure, oraopt = None, 
 | 
| 105 | new_axioms = Symtab.null, parents = []}; | |
| 1526 | 106 | |
| 107 | val cpure_thy = | |
| 1539 | 108 |   Theory {sign = Sign.cpure, oraopt = None, 
 | 
| 109 | new_axioms = Symtab.null, parents = []}; | |
| 1526 | 110 | |
| 111 | ||
| 112 | ||
| 113 | (** extend theory **) | |
| 114 | ||
| 115 | fun err_dup_axms names = | |
| 116 |   error ("Duplicate axiom name(s) " ^ commas_quote names);
 | |
| 117 | ||
| 1539 | 118 | fun ext_thy (thy as Theory {sign, oraopt, new_axioms, parents}) 
 | 
| 119 | sign1 new_axms = | |
| 1526 | 120 | let | 
| 121 | val draft = Sign.is_draft sign; | |
| 122 | val new_axioms1 = | |
| 123 | Symtab.extend_new (if draft then new_axioms else Symtab.null, new_axms) | |
| 124 | handle Symtab.DUPS names => err_dup_axms names; | |
| 125 | val parents1 = if draft then parents else [thy]; | |
| 126 | in | |
| 1539 | 127 |     Theory {sign = sign1, oraopt = oraopt, 
 | 
| 128 | new_axioms = new_axioms1, parents = parents1} | |
| 1526 | 129 | end; | 
| 130 | ||
| 131 | ||
| 132 | (* extend signature of a theory *) | |
| 133 | ||
| 134 | fun ext_sg extfun decls (thy as Theory {sign, ...}) =
 | |
| 135 | ext_thy thy (extfun decls sign) []; | |
| 136 | ||
| 2206 | 137 | val add_classes = ext_sg Sign.add_classes; | 
| 138 | val add_classrel = ext_sg Sign.add_classrel; | |
| 139 | val add_defsort = ext_sg Sign.add_defsort; | |
| 140 | val add_types = ext_sg Sign.add_types; | |
| 141 | val add_tyabbrs = ext_sg Sign.add_tyabbrs; | |
| 142 | val add_tyabbrs_i = ext_sg Sign.add_tyabbrs_i; | |
| 143 | val add_arities = ext_sg Sign.add_arities; | |
| 144 | val add_consts = ext_sg Sign.add_consts; | |
| 145 | val add_consts_i = ext_sg Sign.add_consts_i; | |
| 146 | val add_syntax = ext_sg Sign.add_syntax; | |
| 147 | val add_syntax_i = ext_sg Sign.add_syntax_i; | |
| 148 | val add_modesyntax = curry (ext_sg Sign.add_modesyntax); | |
| 149 | val add_modesyntax_i = curry (ext_sg Sign.add_modesyntax_i); | |
| 150 | val add_trfuns = ext_sg Sign.add_trfuns; | |
| 2385 | 151 | val add_trfunsT = ext_sg Sign.add_trfunsT; | 
| 2206 | 152 | val add_trrules = ext_sg Sign.add_trrules; | 
| 153 | val add_trrules_i = ext_sg Sign.add_trrules_i; | |
| 154 | val add_thyname = ext_sg Sign.add_name; | |
| 1526 | 155 | |
| 156 | ||
| 157 | (* prepare axioms *) | |
| 158 | ||
| 159 | fun err_in_axm name = | |
| 160 |   error ("The error(s) above occurred in axiom " ^ quote name);
 | |
| 161 | ||
| 162 | fun no_vars tm = | |
| 163 | if null (term_vars tm) andalso null (term_tvars tm) then tm | |
| 164 | else error "Illegal schematic variable(s) in term"; | |
| 165 | ||
| 166 | fun cert_axm sg (name, raw_tm) = | |
| 167 | let | |
| 168 | val (t, T, _) = Sign.certify_term sg raw_tm | |
| 169 | handle TYPE arg => error (Sign.exn_type_msg sg arg) | |
| 170 | | TERM (msg, _) => error msg; | |
| 171 | in | |
| 172 | assert (T = propT) "Term not of type prop"; | |
| 173 | (name, no_vars t) | |
| 174 | end | |
| 175 | handle ERROR => err_in_axm name; | |
| 176 | ||
| 177 | (*Some duplication of code with read_def_cterm*) | |
| 178 | fun read_axm sg (name, str) = | |
| 179 | let val ts = Syntax.read (#syn (Sign.rep_sg sg)) propT str; | |
| 180 | val (_, t, _) = | |
| 181 | Sign.infer_types sg (K None) (K None) [] true (ts,propT); | |
| 1960 
ae390b599213
Improved error handling: if there are syntax or type-checking
 paulson parents: 
1539diff
changeset | 182 | in cert_axm sg (name,t) end | 
| 
ae390b599213
Improved error handling: if there are syntax or type-checking
 paulson parents: 
1539diff
changeset | 183 | handle ERROR => err_in_axm name; | 
| 1526 | 184 | |
| 185 | fun inferT_axm sg (name, pre_tm) = | |
| 186 | let val t = #2(Sign.infer_types sg (K None) (K None) [] true | |
| 187 | ([pre_tm], propT)) | |
| 188 | in (name, no_vars t) end | |
| 189 | handle ERROR => err_in_axm name; | |
| 190 | ||
| 191 | ||
| 192 | (* extend axioms of a theory *) | |
| 193 | ||
| 194 | fun ext_axms prep_axm axms (thy as Theory {sign, ...}) =
 | |
| 195 | let | |
| 196 | val sign1 = Sign.make_draft sign; | |
| 197 | val axioms = map (apsnd (Term.compress_term o Logic.varify) o | |
| 198 | prep_axm sign) | |
| 199 | axms; | |
| 200 | in | |
| 201 | ext_thy thy sign1 axioms | |
| 202 | end; | |
| 203 | ||
| 204 | val add_axioms = ext_axms read_axm; | |
| 205 | val add_axioms_i = ext_axms cert_axm; | |
| 206 | ||
| 207 | ||
| 1539 | 208 | (** Set oracle of theory **) | 
| 209 | ||
| 210 | fun set_oracle oracle | |
| 211 |                (thy as Theory {sign, oraopt = None, new_axioms, parents}) =
 | |
| 212 | if Sign.is_draft sign then | |
| 213 | 	Theory {sign = sign, 
 | |
| 214 | oraopt = Some oracle, | |
| 215 | new_axioms = new_axioms, | |
| 216 | parents = parents} | |
| 217 |       else raise THEORY ("Can only set oracle of a draft", [thy])
 | |
| 218 |   | set_oracle _ thy = raise THEORY ("Oracle already set", [thy]);
 | |
| 219 | ||
| 1526 | 220 | |
| 221 | (** merge theories **) | |
| 222 | ||
| 223 | fun merge_thy_list mk_draft thys = | |
| 224 | let | |
| 225 | fun is_union thy = forall (fn t => subthy (t, thy)) thys; | |
| 226 | val is_draft = Sign.is_draft o sign_of; | |
| 227 | ||
| 228 |     fun add_sign (sg, Theory {sign, ...}) =
 | |
| 229 | Sign.merge (sg, sign) handle TERM (msg, _) => error msg; | |
| 230 | in | |
| 1539 | 231 | case (find_first is_union thys, exists is_draft thys) of | 
| 1526 | 232 | (Some thy, _) => thy | 
| 233 |     | (None, true) => raise THEORY ("Illegal merge of draft theories", thys)
 | |
| 234 |     | (None, false) => Theory {
 | |
| 235 | sign = | |
| 236 | (if mk_draft then Sign.make_draft else I) | |
| 237 | (foldl add_sign (Sign.proto_pure, thys)), | |
| 1539 | 238 | oraopt = None, | 
| 1526 | 239 | new_axioms = Symtab.null, | 
| 1539 | 240 | parents = thys} | 
| 1526 | 241 | end; | 
| 242 | ||
| 243 | fun merge_theories (thy1, thy2) = merge_thy_list false [thy1, thy2]; | |
| 244 | ||
| 245 | ||
| 246 | end; | |
| 247 | ||
| 248 | open Theory; |