| author | wenzelm | 
| Sun, 11 Nov 2001 21:31:52 +0100 | |
| changeset 12139 | d51d50636332 | 
| parent 12130 | 30d9143aff7e | 
| child 12147 | 64e69a8a945f | 
| permissions | -rw-r--r-- | 
| 5819 | 1 | (* Title: Pure/Isar/proof_context.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 8807 | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 5819 | 5 | |
| 6 | Proof context information. | |
| 7 | *) | |
| 8 | ||
| 9 | signature PROOF_CONTEXT = | |
| 10 | sig | |
| 11 | type context | |
| 9470 | 12 | type exporter | 
| 5819 | 13 | exception CONTEXT of string * context | 
| 14 | val theory_of: context -> theory | |
| 15 | val sign_of: context -> Sign.sg | |
| 12093 | 16 | val syntax_of: context -> Syntax.syntax * string list * string list | 
| 12057 | 17 | val fixed_names_of: context -> string list | 
| 18 | val assumptions_of: context -> (cterm list * exporter) list | |
| 7557 | 19 | val prems_of: context -> thm list | 
| 5819 | 20 | val print_proof_data: theory -> unit | 
| 5874 | 21 | val init: theory -> context | 
| 12057 | 22 | val is_fixed: context -> string -> bool | 
| 5819 | 23 | val read_typ: context -> string -> typ | 
| 9504 | 24 | val read_typ_no_norm: context -> string -> typ | 
| 5819 | 25 | val cert_typ: context -> typ -> typ | 
| 9504 | 26 | val cert_typ_no_norm: context -> typ -> typ | 
| 10583 | 27 | val get_skolem: context -> string -> string | 
| 9291 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 28 | val intern_skolem: context -> term -> term | 
| 9133 | 29 | val extern_skolem: context -> term -> term | 
| 5874 | 30 | val read_termTs: context -> (string * typ) list -> term list * (indexname * typ) list | 
| 5819 | 31 | val read_term: context -> string -> term | 
| 32 | val read_prop: context -> string -> term | |
| 11925 | 33 | val read_prop_schematic: context -> string -> term | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 34 | val read_terms: context -> string list -> term list | 
| 8096 | 35 | val read_termT_pats: context -> (string * typ) list -> term list | 
| 36 | val read_term_pats: typ -> context -> string list -> term list | |
| 37 | val read_prop_pats: context -> string list -> term list | |
| 5819 | 38 | val cert_term: context -> term -> term | 
| 39 | val cert_prop: context -> term -> term | |
| 8096 | 40 | val cert_term_pats: typ -> context -> term list -> term list | 
| 41 | val cert_prop_pats: context -> term list -> term list | |
| 5819 | 42 | val declare_term: term -> context -> context | 
| 43 | val declare_terms: term list -> context -> context | |
| 7925 | 44 | val warn_extra_tfrees: context -> context -> context | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 45 | val generalizeT: context -> context -> typ -> typ | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 46 | val generalize: context -> context -> term -> term | 
| 9470 | 47 | val find_free: term -> string -> term option | 
| 12008 | 48 | val export: bool -> context -> context -> thm -> thm Seq.seq | 
| 10810 | 49 | val drop_schematic: indexname * term option -> indexname * term option | 
| 50 | val add_binds: (indexname * string option) list -> context -> context | |
| 51 | val add_binds_i: (indexname * term option) list -> context -> context | |
| 7925 | 52 | val auto_bind_goal: term -> context -> context | 
| 53 | val auto_bind_facts: string -> term list -> context -> context | |
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 54 | val match_bind: bool -> (string list * string) list -> context -> context | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 55 | val match_bind_i: bool -> (term list * term) list -> context -> context | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 56 | val read_propp: context * (string * (string list * string list)) list list | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 57 | -> context * (term * (term list * term list)) list list | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 58 | val cert_propp: context * (term * (term list * term list)) list list | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 59 | -> context * (term * (term list * term list)) list list | 
| 10554 | 60 | val read_propp_schematic: context * (string * (string list * string list)) list list | 
| 61 | -> context * (term * (term list * term list)) list list | |
| 62 | val cert_propp_schematic: context * (term * (term list * term list)) list list | |
| 63 | -> context * (term * (term list * term list)) list list | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 64 | val bind_propp: context * (string * (string list * string list)) list list | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 65 | -> context * (term list list * (context -> context)) | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 66 | val bind_propp_i: context * (term * (term list * term list)) list list | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 67 | -> context * (term list list * (context -> context)) | 
| 10554 | 68 | val bind_propp_schematic: context * (string * (string list * string list)) list list | 
| 69 | -> context * (term list list * (context -> context)) | |
| 70 | val bind_propp_schematic_i: context * (term * (term list * term list)) list list | |
| 71 | -> context * (term list list * (context -> context)) | |
| 6091 | 72 | val get_thm: context -> string -> thm | 
| 9566 | 73 | val get_thm_closure: context -> string -> thm | 
| 6091 | 74 | val get_thms: context -> string -> thm list | 
| 9566 | 75 | val get_thms_closure: context -> string -> thm list | 
| 6091 | 76 | val put_thm: string * thm -> context -> context | 
| 77 | val put_thms: string * thm list -> context -> context | |
| 78 | val put_thmss: (string * thm list) list -> context -> context | |
| 7606 | 79 | val reset_thms: string -> context -> context | 
| 9196 | 80 | val have_thmss: | 
| 81 | ((string * context attribute list) * (thm list * context attribute list) list) list -> | |
| 82 | context -> context * (string * thm list) list | |
| 11918 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 83 | val export_assume: exporter | 
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 84 | val export_presume: exporter | 
| 12086 | 85 | val cert_def: context -> term -> string * term | 
| 11918 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 86 | val export_def: exporter | 
| 9470 | 87 | val assume: exporter | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 88 | -> ((string * context attribute list) * (string * (string list * string list)) list) list | 
| 11925 | 89 | -> context -> context * (string * thm list) list | 
| 9470 | 90 | val assume_i: exporter | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 91 | -> ((string * context attribute list) * (term * (term list * term list)) list) list | 
| 11925 | 92 | -> context -> context * (string * thm list) list | 
| 8096 | 93 | val read_vars: context * (string list * string option) -> context * (string list * typ option) | 
| 94 | val cert_vars: context * (string list * typ option) -> context * (string list * typ option) | |
| 7411 | 95 | val fix: (string list * string option) list -> context -> context | 
| 7663 | 96 | val fix_i: (string list * typ option) list -> context -> context | 
| 11925 | 97 | val fix_direct: (string list * typ option) list -> context -> context | 
| 12016 | 98 | val fix_frees: term list -> context -> context | 
| 9291 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 99 | val bind_skolem: context -> string list -> term -> term | 
| 11793 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 100 | val get_case: context -> string -> string option list -> RuleCases.T | 
| 8373 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 wenzelm parents: 
8186diff
changeset | 101 | val add_cases: (string * RuleCases.T) list -> context -> context | 
| 10830 | 102 | val apply_case: RuleCases.T -> context -> context * ((indexname * term option) list * term list) | 
| 10810 | 103 | val show_hyps: bool ref | 
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 104 | val add_syntax: (string * typ option * mixfix option) list -> context -> context | 
| 12057 | 105 | val pretty_term: context -> term -> Pretty.T | 
| 106 | val pretty_typ: context -> typ -> Pretty.T | |
| 107 | val pretty_sort: context -> sort -> Pretty.T | |
| 108 | val pretty_thm: context -> thm -> Pretty.T | |
| 109 | val pretty_thms: context -> thm list -> Pretty.T | |
| 110 | val string_of_term: context -> term -> string | |
| 10810 | 111 | val verbose: bool ref | 
| 112 |   val setmp_verbose: ('a -> 'b) -> 'a -> 'b
 | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 113 | val print_syntax: context -> unit | 
| 10810 | 114 | val print_binds: context -> unit | 
| 12057 | 115 | val print_lthms: context -> unit | 
| 10810 | 116 | val print_cases: context -> unit | 
| 117 | val prems_limit: int ref | |
| 12057 | 118 | val pretty_asms: context -> Pretty.T list | 
| 10810 | 119 | val pretty_context: context -> Pretty.T list | 
| 5819 | 120 | val setup: (theory -> theory) list | 
| 121 | end; | |
| 122 | ||
| 8151 | 123 | signature PRIVATE_PROOF_CONTEXT = | 
| 5819 | 124 | sig | 
| 125 | include PROOF_CONTEXT | |
| 126 | val init_data: Object.kind -> (theory -> Object.T) * (context -> Object.T -> unit) | |
| 127 | -> theory -> theory | |
| 128 | val print_data: Object.kind -> context -> unit | |
| 129 | val get_data: Object.kind -> (Object.T -> 'a) -> context -> 'a | |
| 130 |   val put_data: Object.kind -> ('a -> Object.T) -> 'a -> context -> context
 | |
| 131 | end; | |
| 132 | ||
| 8151 | 133 | structure ProofContext: PRIVATE_PROOF_CONTEXT = | 
| 5819 | 134 | struct | 
| 135 | ||
| 136 | ||
| 137 | (** datatype context **) | |
| 138 | ||
| 11816 | 139 | type exporter = bool -> cterm list -> thm -> thm Seq.seq; | 
| 9470 | 140 | |
| 5819 | 141 | datatype context = | 
| 142 | Context of | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 143 |    {thy: theory,                                           (*current theory*)
 | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 144 | syntax: Syntax.syntax * string list * string list, (*syntax with structs and mixfixed*) | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 145 | data: Object.T Symtab.table, (*user data*) | 
| 5819 | 146 | asms: | 
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 147 | ((cterm list * exporter) list * (*assumes: A ==> _*) | 
| 6931 | 148 | (string * thm list) list) * | 
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 149 | ((string * string) list * string list), (*fixes: !!x. _*) | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 150 | binds: (term * typ) option Vartab.table, (*term bindings*) | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 151 | thms: thm list option Symtab.table, (*local thms*) | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 152 | cases: (string * RuleCases.T) list, (*local contexts*) | 
| 5819 | 153 | defs: | 
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 154 | typ Vartab.table * (*type constraints*) | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 155 | sort Vartab.table * (*default sorts*) | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 156 | (string list * term list Symtab.table)}; (*used type variables*) | 
| 5819 | 157 | |
| 158 | exception CONTEXT of string * context; | |
| 159 | ||
| 160 | ||
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 161 | fun make_context (thy, syntax, data, asms, binds, thms, cases, defs) = | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 162 |   Context {thy = thy, syntax = syntax, data = data, asms = asms, binds = binds,
 | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 163 | thms = thms, cases = cases, defs = defs}; | 
| 5819 | 164 | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 165 | fun map_context f (Context {thy, syntax, data, asms, binds, thms, cases, defs}) =
 | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 166 | make_context (f (thy, syntax, data, asms, binds, thms, cases, defs)); | 
| 5819 | 167 | |
| 168 | fun theory_of (Context {thy, ...}) = thy;
 | |
| 169 | val sign_of = Theory.sign_of o theory_of; | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 170 | fun syntax_of (Context {syntax, ...}) = syntax;
 | 
| 5819 | 171 | |
| 12130 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 172 | fun vars_of (Context {asms = (_, vars), ...}) = vars;
 | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 173 | val fixes_of = #1 o vars_of; | 
| 12057 | 174 | val fixed_names_of = map #2 o fixes_of; | 
| 175 | fun is_fixed (Context {asms = (_, (fixes, _)), ...}) x = exists (equal x o #2) fixes;
 | |
| 176 | fun used_table (Context {defs = (_, _, (_, tab)), ...}) = tab;
 | |
| 177 | ||
| 178 | fun assumptions_of (Context {asms = ((asms, _), _), ...}) = asms;
 | |
| 179 | fun prems_of (Context {asms = ((_, prems), _), ...}) = flat (map #2 prems);
 | |
| 180 | ||
| 7270 | 181 | |
| 5819 | 182 | |
| 183 | (** user data **) | |
| 184 | ||
| 185 | (* errors *) | |
| 186 | ||
| 187 | fun of_theory thy = "\nof theory " ^ Sign.str_of_sg (Theory.sign_of thy); | |
| 188 | ||
| 189 | fun err_inconsistent kinds = | |
| 190 |   error ("Attempt to merge different versions of " ^ commas_quote kinds ^ " proof data");
 | |
| 191 | ||
| 192 | fun err_dup_init thy kind = | |
| 193 |   error ("Duplicate initialization of " ^ quote kind ^ " proof data" ^ of_theory thy);
 | |
| 194 | ||
| 195 | fun err_undef ctxt kind = | |
| 196 |   raise CONTEXT ("Tried to access undefined " ^ quote kind ^ " proof data", ctxt);
 | |
| 197 | ||
| 198 | fun err_uninit ctxt kind = | |
| 199 |   raise CONTEXT ("Tried to access uninitialized " ^ quote kind ^ " proof data" ^
 | |
| 200 | of_theory (theory_of ctxt), ctxt); | |
| 201 | ||
| 202 | fun err_access ctxt kind = | |
| 203 |   raise CONTEXT ("Unauthorized access to " ^ quote kind ^ " proof data" ^
 | |
| 204 | of_theory (theory_of ctxt), ctxt); | |
| 205 | ||
| 206 | ||
| 207 | (* data kind 'Isar/proof_data' *) | |
| 208 | ||
| 209 | structure ProofDataDataArgs = | |
| 210 | struct | |
| 211 | val name = "Isar/proof_data"; | |
| 212 | type T = (Object.kind * ((theory -> Object.T) * (context -> Object.T -> unit))) Symtab.table; | |
| 213 | ||
| 214 | val empty = Symtab.empty; | |
| 6550 | 215 | val copy = I; | 
| 12123 | 216 | val finish = I; | 
| 5819 | 217 | val prep_ext = I; | 
| 218 | fun merge tabs = Symtab.merge (Object.eq_kind o pairself fst) tabs | |
| 219 | handle Symtab.DUPS kinds => err_inconsistent kinds; | |
| 220 | fun print _ tab = Pretty.writeln (Pretty.strs (map #1 (Symtab.dest tab))); | |
| 221 | end; | |
| 222 | ||
| 223 | structure ProofDataData = TheoryDataFun(ProofDataDataArgs); | |
| 224 | val print_proof_data = ProofDataData.print; | |
| 225 | ||
| 226 | ||
| 227 | (* init proof data *) | |
| 228 | ||
| 229 | fun init_data kind meths thy = | |
| 230 | let | |
| 231 | val name = Object.name_of_kind kind; | |
| 232 | val tab = Symtab.update_new ((name, (kind, meths)), ProofDataData.get thy) | |
| 233 | handle Symtab.DUP _ => err_dup_init thy name; | |
| 234 | in thy |> ProofDataData.put tab end; | |
| 235 | ||
| 236 | ||
| 237 | (* access data *) | |
| 238 | ||
| 239 | fun lookup_data (ctxt as Context {data, ...}) kind =
 | |
| 240 | let | |
| 241 | val thy = theory_of ctxt; | |
| 242 | val name = Object.name_of_kind kind; | |
| 243 | in | |
| 244 | (case Symtab.lookup (ProofDataData.get thy, name) of | |
| 245 | Some (k, meths) => | |
| 246 | if Object.eq_kind (kind, k) then | |
| 247 | (case Symtab.lookup (data, name) of | |
| 248 | Some x => (x, meths) | |
| 249 | | None => err_undef ctxt name) | |
| 250 | else err_access ctxt name | |
| 251 | | None => err_uninit ctxt name) | |
| 252 | end; | |
| 253 | ||
| 254 | fun get_data kind f ctxt = | |
| 255 | let val (x, _) = lookup_data ctxt kind | |
| 256 | in f x handle Match => Object.kind_error kind end; | |
| 257 | ||
| 258 | fun print_data kind ctxt = | |
| 259 | let val (x, (_, prt)) = lookup_data ctxt kind | |
| 260 | in prt ctxt x end; | |
| 261 | ||
| 262 | fun put_data kind f x ctxt = | |
| 8373 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 wenzelm parents: 
8186diff
changeset | 263 | (lookup_data ctxt kind; | 
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 264 | map_context (fn (thy, syntax, data, asms, binds, thms, cases, defs) => | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 265 | (thy, syntax, Symtab.update ((Object.name_of_kind kind, f x), data), | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 266 | asms, binds, thms, cases, defs)) ctxt); | 
| 5819 | 267 | |
| 268 | ||
| 269 | (* init context *) | |
| 270 | ||
| 5874 | 271 | fun init thy = | 
| 272 | let val data = Symtab.map (fn (_, (f, _)) => f thy) (ProofDataData.get thy) in | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 273 | make_context (thy, (Theory.syn_of thy, [], []), data, (([], []), ([], [])), Vartab.empty, | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 274 | Symtab.empty, [], (Vartab.empty, Vartab.empty, ([], Symtab.empty))) | 
| 5819 | 275 | end; | 
| 276 | ||
| 277 | ||
| 8096 | 278 | |
| 12093 | 279 | (** local syntax **) | 
| 280 | ||
| 281 | val fixedN = "\\<^fixed>"; | |
| 282 | val structN = "\\<^struct>"; | |
| 283 | ||
| 12100 | 284 | fun the_struct structs i = | 
| 285 | if 1 <= i andalso i <= length structs then Library.nth_elem (i - 1, structs) | |
| 286 |   else raise ERROR_MESSAGE ("Illegal reference to implicit structure #" ^ string_of_int i);
 | |
| 287 | ||
| 288 | ||
| 289 | (* print (ast) translations *) | |
| 290 | ||
| 291 | fun index_tr' 1 = Syntax.const "_noindex" | |
| 292 | | index_tr' i = Syntax.const "_index" $ Syntax.const (Library.string_of_int i); | |
| 293 | ||
| 294 | fun context_tr' ctxt = | |
| 295 | let | |
| 296 | val (_, structs, mixfixed) = syntax_of ctxt; | |
| 297 | ||
| 298 | fun tr' (t $ u) = tr' t $ tr' u | |
| 299 | | tr' (Abs (x, T, t)) = Abs (x, T, tr' t) | |
| 300 | | tr' (t as Free (x, T)) = | |
| 301 | let val i = Library.find_index (equal x) structs + 1 in | |
| 302 | if 1 <= i andalso i <= 9 then Syntax.const "_struct" $ index_tr' i | |
| 303 | else if x mem_string mixfixed then Const (fixedN ^ x, T) | |
| 304 | else t | |
| 305 | end | |
| 306 | | tr' a = a; | |
| 307 | in tr' end; | |
| 308 | ||
| 309 | fun index_ast_tr' structs s = | |
| 310 | (case Syntax.read_nat s of | |
| 311 | Some i => Syntax.Variable (the_struct structs i handle ERROR_MESSAGE _ => raise Match) | |
| 312 | | None => raise Match); | |
| 313 | ||
| 314 | fun struct_ast_tr' structs [Syntax.Constant "_noindex"] = | |
| 315 | index_ast_tr' structs "1" | |
| 316 | | struct_ast_tr' structs [Syntax.Appl [Syntax.Constant "_index", Syntax.Constant s]] = | |
| 317 | index_ast_tr' structs s | |
| 318 | | struct_ast_tr' _ _ = raise Match; | |
| 319 | ||
| 320 | ||
| 321 | (* parse translations *) | |
| 322 | ||
| 323 | fun fixed_tr x = (fixedN ^ x, curry Term.list_comb (Syntax.free x)); | |
| 324 | ||
| 325 | fun index_tr (Const ("_noindex", _)) = 1
 | |
| 326 |   | index_tr (t as (Const ("_index", _) $ Const (s, _))) =
 | |
| 327 |       (case Syntax.read_nat s of Some n => n | None => raise TERM ("index_tr", [t]))
 | |
| 328 |   | index_tr t = raise TERM ("index_tr", [t]);
 | |
| 329 | ||
| 330 | fun struct_tr structs (idx :: ts) = Syntax.free (the_struct structs (index_tr idx)) | |
| 331 |   | struct_tr _ ts = raise TERM ("struct_tr", ts);
 | |
| 332 | ||
| 12093 | 333 | |
| 334 | (* add_syntax and syn_of *) | |
| 335 | ||
| 336 | local | |
| 337 | ||
| 338 | val aT = TFree ("'a", logicS);
 | |
| 339 | ||
| 340 | fun mixfix x None mx = (fixedN ^ x, replicate (Syntax.mixfix_args mx) aT ---> aT, mx) | |
| 341 | | mixfix x (Some T) mx = (fixedN ^ x, T, mx); | |
| 342 | ||
| 343 | fun prep_mixfix (_, _, None) = None | |
| 344 | | prep_mixfix (x, opt_T, Some mx) = Some (mixfix x opt_T mx); | |
| 345 | ||
| 346 | fun prep_mixfix' (_, _, None) = None | |
| 347 | | prep_mixfix' (x, _, Some Syntax.NoSyn) = None | |
| 348 | | prep_mixfix' (x, opt_T, _) = Some (x, mixfix x opt_T (Syntax.Delimfix x)); | |
| 349 | ||
| 350 | fun prep_struct (x, _, None) = Some x | |
| 351 | | prep_struct _ = None; | |
| 352 | ||
| 353 | in | |
| 354 | ||
| 355 | fun add_syntax decls = | |
| 356 | map_context (fn (thy, (syn, structs, mixfixed), data, asms, binds, thms, cases, defs) => | |
| 357 | let | |
| 358 | val structs' = structs @ mapfilter prep_struct decls; | |
| 359 | val mxs = mapfilter prep_mixfix decls; | |
| 360 | val (fixed, mxs_output) = Library.split_list (mapfilter prep_mixfix' decls); | |
| 361 | val trs = map fixed_tr fixed; | |
| 362 | val syn' = syn | |
| 363 |         |> Syntax.extend_const_gram ("", false) mxs_output
 | |
| 364 |         |> Syntax.extend_const_gram ("", true) mxs
 | |
| 365 | |> Syntax.extend_trfuns ([], trs, [], []); | |
| 366 | in (thy, (syn', structs', fixed @ mixfixed), data, asms, binds, thms, cases, defs) end) | |
| 367 | ||
| 368 | fun syn_of (Context {syntax = (syn, structs, _), ...}) =
 | |
| 12100 | 369 | syn |> Syntax.extend_trfuns | 
| 370 |     ([], [("_struct", struct_tr structs)], [], [("_struct", struct_ast_tr' structs)]);
 | |
| 12093 | 371 | |
| 372 | end; | |
| 373 | ||
| 374 | ||
| 375 | ||
| 7663 | 376 | (** default sorts and types **) | 
| 377 | ||
| 7670 | 378 | fun def_sort (Context {defs = (_, sorts, _), ...}) xi = Vartab.lookup (sorts, xi);
 | 
| 7663 | 379 | |
| 7670 | 380 | fun def_type (Context {binds, defs = (types, _, _), ...}) is_pat xi =
 | 
| 7663 | 381 | (case Vartab.lookup (types, xi) of | 
| 382 | None => | |
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 383 | if is_pat then None else | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 384 | (case Vartab.lookup (binds, xi) of | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 385 | Some (Some (_, T)) => Some (TypeInfer.polymorphicT T) | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 386 | | _ => None) | 
| 7663 | 387 | | some => some); | 
| 388 | ||
| 389 | ||
| 5819 | 390 | |
| 391 | (** prepare types **) | |
| 392 | ||
| 9504 | 393 | local | 
| 394 | ||
| 395 | fun read_typ_aux read ctxt s = | |
| 12093 | 396 | transform_error (read (syn_of ctxt) (sign_of ctxt, def_sort ctxt)) s | 
| 7663 | 397 | handle ERROR_MESSAGE msg => raise CONTEXT (msg, ctxt); | 
| 5819 | 398 | |
| 10554 | 399 | fun cert_typ_aux cert ctxt raw_T = | 
| 400 | cert (sign_of ctxt) raw_T | |
| 401 | handle TYPE (msg, _, _) => raise CONTEXT (msg, ctxt); | |
| 9504 | 402 | |
| 403 | in | |
| 404 | ||
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 405 | val read_typ = read_typ_aux Sign.read_typ'; | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 406 | val read_typ_no_norm = read_typ_aux Sign.read_typ_no_norm'; | 
| 11925 | 407 | val cert_typ = cert_typ_aux Sign.certify_typ; | 
| 9504 | 408 | val cert_typ_no_norm = cert_typ_aux Sign.certify_typ_no_norm; | 
| 409 | ||
| 410 | end; | |
| 411 | ||
| 5819 | 412 | |
| 7679 | 413 | (* internalize Skolem constants *) | 
| 414 | ||
| 10583 | 415 | fun lookup_skolem ctxt x = assoc (fixes_of ctxt, x); | 
| 416 | fun get_skolem ctxt x = if_none (lookup_skolem ctxt x) x; | |
| 7679 | 417 | |
| 9325 | 418 | fun no_skolem no_internal ctxt x = | 
| 9291 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 419 | if can Syntax.dest_skolem x then | 
| 7679 | 420 |     raise CONTEXT ("Illegal reference to internal Skolem constant: " ^ quote x, ctxt)
 | 
| 9325 | 421 | else if no_internal andalso can Syntax.dest_internal x then | 
| 9291 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 422 |     raise CONTEXT ("Illegal reference to internal variable: " ^ quote x, ctxt)
 | 
| 7679 | 423 | else x; | 
| 424 | ||
| 9291 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 425 | fun intern_skolem ctxt = | 
| 7679 | 426 | let | 
| 427 | fun intern (t as Free (x, T)) = | |
| 10583 | 428 | (case lookup_skolem ctxt (no_skolem true ctxt x) of | 
| 7679 | 429 | Some x' => Free (x', T) | 
| 430 | | None => t) | |
| 431 | | intern (t $ u) = intern t $ intern u | |
| 432 | | intern (Abs (x, T, t)) = Abs (x, T, intern t) | |
| 433 | | intern a = a; | |
| 434 | in intern end; | |
| 435 | ||
| 436 | ||
| 9133 | 437 | (* externalize Skolem constants -- for printing purposes only *) | 
| 438 | ||
| 439 | fun extern_skolem ctxt = | |
| 440 | let | |
| 441 | val rev_fixes = map Library.swap (fixes_of ctxt); | |
| 442 | ||
| 443 | fun extern (t as Free (x, T)) = | |
| 444 | (case assoc (rev_fixes, x) of | |
| 10583 | 445 | Some x' => Free (if lookup_skolem ctxt x' = Some x then x' else NameSpace.hidden x', T) | 
| 9133 | 446 | | None => t) | 
| 447 | | extern (t $ u) = extern t $ extern u | |
| 448 | | extern (Abs (x, T, t)) = Abs (x, T, extern t) | |
| 449 | | extern a = a; | |
| 450 | in extern end | |
| 451 | ||
| 8096 | 452 | |
| 5819 | 453 | (** prepare terms and propositions **) | 
| 454 | ||
| 455 | (* | |
| 456 | (1) read / certify wrt. signature of context | |
| 457 | (2) intern Skolem constants | |
| 458 | (3) expand term bindings | |
| 459 | *) | |
| 460 | ||
| 461 | ||
| 462 | (* read / certify wrt. signature *) (*exception ERROR*) (*exception TERM*) | |
| 463 | ||
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 464 | fun read_def_termTs freeze syn sg (types, sorts, used) sTs = | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 465 | Sign.read_def_terms' syn (sg, types, sorts) used freeze sTs; | 
| 5874 | 466 | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 467 | fun read_def_termT freeze syn sg defs sT = apfst hd (read_def_termTs freeze syn sg defs [sT]); | 
| 5874 | 468 | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 469 | fun read_term_sg freeze syn sg defs s = | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 470 | #1 (read_def_termT freeze syn sg defs (s, TypeInfer.logicT)); | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 471 | |
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 472 | fun read_prop_sg freeze syn sg defs s = #1 (read_def_termT freeze syn sg defs (s, propT)); | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 473 | |
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 474 | fun read_terms_sg freeze syn sg defs = | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 475 | #1 o read_def_termTs freeze syn sg defs o map (rpair TypeInfer.logicT); | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 476 | |
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 477 | fun read_props_sg freeze syn sg defs = #1 o read_def_termTs freeze syn sg defs o map (rpair propT); | 
| 5819 | 478 | |
| 479 | ||
| 480 | fun cert_term_sg sg t = Thm.term_of (Thm.cterm_of sg t); | |
| 481 | ||
| 482 | fun cert_prop_sg sg tm = | |
| 483 | let | |
| 484 | val ctm = Thm.cterm_of sg tm; | |
| 485 |     val {t, T, ...} = Thm.rep_cterm ctm;
 | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 486 |   in if T = propT then t else raise TERM ("Term not of type prop", [t]) end;
 | 
| 5819 | 487 | |
| 488 | ||
| 489 | (* norm_term *) | |
| 490 | ||
| 491 | (*beta normal form for terms (not eta normal form), chase variables in | |
| 492 | bindings environment (code taken from Pure/envir.ML)*) | |
| 493 | ||
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 494 | fun unifyT ctxt (T, U) = | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 495 | let val maxidx = Int.max (Term.maxidx_of_typ T, Term.maxidx_of_typ U) | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 496 | in #1 (Type.unify (Sign.tsig_of (sign_of ctxt)) maxidx Vartab.empty (T, U)) end; | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 497 | |
| 10554 | 498 | fun norm_term (ctxt as Context {binds, ...}) schematic =
 | 
| 5819 | 499 | let | 
| 500 | (*raised when norm has no effect on a term, to do sharing instead of copying*) | |
| 501 | exception SAME; | |
| 502 | ||
| 503 | fun norm (t as Var (xi, T)) = | |
| 504 | (case Vartab.lookup (binds, xi) of | |
| 7606 | 505 | Some (Some (u, U)) => | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 506 | let | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 507 | val env = unifyT ctxt (T, U) handle Type.TUNIFY => | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 508 |                   raise TYPE ("norm_term: ill-typed variable assigment", [T, U], [t, u]);
 | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 509 | val u' = Term.subst_TVars_Vartab env u; | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 510 | in norm u' handle SAME => u' end | 
| 10554 | 511 | | _ => | 
| 512 | if schematic then raise SAME | |
| 513 |             else raise CONTEXT ("Unbound schematic variable: " ^ Syntax.string_of_vname xi, ctxt))
 | |
| 5819 | 514 | | norm (Abs (a, T, body)) = Abs (a, T, norm body) | 
| 515 | | norm (Abs (_, _, body) $ t) = normh (subst_bound (t, body)) | |
| 516 | | norm (f $ t) = | |
| 517 | ((case norm f of | |
| 518 | Abs (_, _, body) => normh (subst_bound (t, body)) | |
| 519 | | nf => nf $ (norm t handle SAME => t)) handle SAME => f $ norm t) | |
| 520 | | norm _ = raise SAME | |
| 521 | and normh t = norm t handle SAME => t | |
| 522 | in normh end; | |
| 523 | ||
| 524 | ||
| 6550 | 525 | (* dummy patterns *) | 
| 526 | ||
| 9540 | 527 | fun prepare_dummies t = #2 (Term.replace_dummy_patterns (1, t)); | 
| 6762 | 528 | |
| 9540 | 529 | fun reject_dummies ctxt t = Term.no_dummy_patterns t | 
| 530 |   handle TERM _ => raise CONTEXT ("Illegal dummy pattern(s) in term", ctxt);
 | |
| 6550 | 531 | |
| 532 | ||
| 5819 | 533 | (* read terms *) | 
| 534 | ||
| 10554 | 535 | local | 
| 536 | ||
| 537 | fun gen_read read app is_pat schematic (ctxt as Context {defs = (_, _, (used, _)), ...}) s =
 | |
| 12093 | 538 | (transform_error (read (syn_of ctxt) | 
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 539 | (sign_of ctxt) (def_type ctxt is_pat, def_sort ctxt, used)) s | 
| 7663 | 540 | handle TERM (msg, _) => raise CONTEXT (msg, ctxt) | 
| 541 | | ERROR_MESSAGE msg => raise CONTEXT (msg, ctxt)) | |
| 9291 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 542 | |> app (intern_skolem ctxt) | 
| 10554 | 543 | |> app (if is_pat then I else norm_term ctxt schematic) | 
| 7663 | 544 | |> app (if is_pat then prepare_dummies else (reject_dummies ctxt)); | 
| 5819 | 545 | |
| 10554 | 546 | in | 
| 547 | ||
| 548 | val read_termTs = gen_read (read_def_termTs false) (apfst o map) false false; | |
| 549 | val read_termT_pats = #1 oo gen_read (read_def_termTs false) (apfst o map) true false; | |
| 8096 | 550 | |
| 551 | fun read_term_pats T ctxt pats = read_termT_pats ctxt (map (rpair T) pats); | |
| 552 | val read_prop_pats = read_term_pats propT; | |
| 553 | ||
| 10554 | 554 | val read_term = gen_read (read_term_sg true) I false false; | 
| 555 | val read_prop = gen_read (read_prop_sg true) I false false; | |
| 11925 | 556 | val read_prop_schematic = gen_read (read_prop_sg true) I false true; | 
| 10554 | 557 | val read_terms = gen_read (read_terms_sg true) map false false; | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 558 | val read_props = gen_read (read_props_sg true) map false; | 
| 5819 | 559 | |
| 10554 | 560 | end; | 
| 561 | ||
| 5819 | 562 | |
| 563 | (* certify terms *) | |
| 564 | ||
| 10554 | 565 | local | 
| 566 | ||
| 567 | fun gen_cert cert is_pat schematic ctxt t = t | |
| 568 | |> (if is_pat then I else norm_term ctxt schematic) | |
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 569 | |> (fn t' => cert (sign_of ctxt) t' handle TERM (msg, _) => raise CONTEXT (msg, ctxt)); | 
| 5819 | 570 | |
| 10554 | 571 | in | 
| 8096 | 572 | |
| 10554 | 573 | val cert_term = gen_cert cert_term_sg false false; | 
| 574 | val cert_prop = gen_cert cert_prop_sg false false; | |
| 575 | val cert_props = map oo gen_cert cert_prop_sg false; | |
| 576 | ||
| 577 | fun cert_term_pats _ = map o gen_cert cert_term_sg true false; | |
| 578 | val cert_prop_pats = map o gen_cert cert_prop_sg true false; | |
| 579 | ||
| 580 | end; | |
| 5819 | 581 | |
| 582 | ||
| 583 | (* declare terms *) | |
| 584 | ||
| 10381 | 585 | local | 
| 586 | ||
| 5819 | 587 | val ins_types = foldl_aterms | 
| 588 | (fn (types, Free (x, T)) => Vartab.update (((x, ~1), T), types) | |
| 589 | | (types, Var v) => Vartab.update (v, types) | |
| 590 | | (types, _) => types); | |
| 591 | ||
| 592 | val ins_sorts = foldl_types (foldl_atyps | |
| 593 | (fn (sorts, TFree (x, S)) => Vartab.update (((x, ~1), S), sorts) | |
| 594 | | (sorts, TVar v) => Vartab.update (v, sorts) | |
| 595 | | (sorts, _) => sorts)); | |
| 596 | ||
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 597 | val ins_used = foldl_term_types (fn t => foldl_atyps | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 598 | (fn ((used, tab), TFree (x, _)) => (x ins used, Symtab.update_multi ((x, t), tab)) | 
| 5819 | 599 | | (used, _) => used)); | 
| 600 | ||
| 7663 | 601 | fun ins_skolem def_ty = foldr | 
| 5994 | 602 | (fn ((x, x'), types) => | 
| 7663 | 603 | (case def_ty x' of | 
| 5994 | 604 | Some T => Vartab.update (((x, ~1), T), types) | 
| 605 | | None => types)); | |
| 606 | ||
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 607 | fun map_defaults f = map_context (fn (thy, syntax, data, asms, binds, thms, cases, defs) => | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 608 | (thy, syntax, data, asms, binds, thms, cases, f defs)); | 
| 5819 | 609 | |
| 10381 | 610 | fun declare_syn (ctxt, t) = | 
| 5819 | 611 | ctxt | 
| 7670 | 612 | |> map_defaults (fn (types, sorts, used) => (ins_types (types, t), sorts, used)) | 
| 10381 | 613 | |> map_defaults (fn (types, sorts, used) => (types, ins_sorts (sorts, t), used)); | 
| 614 | ||
| 615 | fun declare_occ (ctxt as Context {asms = (_, (fixes, _)), ...}, t) =
 | |
| 616 | declare_syn (ctxt, t) | |
| 7670 | 617 | |> map_defaults (fn (types, sorts, used) => (types, sorts, ins_used (used, t))) | 
| 618 | |> map_defaults (fn (types, sorts, used) => | |
| 619 | (ins_skolem (fn x => Vartab.lookup (types, (x, ~1))) (fixes, types), sorts, used)); | |
| 5819 | 620 | |
| 10381 | 621 | in | 
| 5819 | 622 | |
| 10381 | 623 | fun declare_term t ctxt = declare_occ (ctxt, t); | 
| 624 | fun declare_terms ts ctxt = foldl declare_occ (ctxt, ts); | |
| 12130 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 625 | fun declare_terms_syntax ts ctxt = foldl declare_syn (ctxt, ts); | 
| 10381 | 626 | |
| 627 | end; | |
| 5819 | 628 | |
| 629 | ||
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 630 | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 631 | (** pretty printing **) | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 632 | |
| 12093 | 633 | fun pretty_term ctxt = Sign.pretty_term' (syn_of ctxt) (sign_of ctxt) o context_tr' ctxt; | 
| 12057 | 634 | val pretty_typ = Sign.pretty_typ o sign_of; | 
| 635 | val pretty_sort = Sign.pretty_sort o sign_of; | |
| 636 | ||
| 637 | val string_of_term = Pretty.string_of oo pretty_term; | |
| 638 | ||
| 639 | val show_hyps = ref false; | |
| 640 | ||
| 641 | fun pretty_thm ctxt thm = | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 642 | if ! show_hyps then setmp Display.show_hyps true (fn () => | 
| 12086 | 643 | Display.pretty_thm_aux (pretty_sort ctxt, pretty_term ctxt) false thm) () | 
| 12057 | 644 | else pretty_term ctxt (#prop (Thm.rep_thm thm)); | 
| 645 | ||
| 646 | fun pretty_thms ctxt [th] = pretty_thm ctxt th | |
| 647 | | pretty_thms ctxt ths = Pretty.blk (0, Pretty.fbreaks (map (pretty_thm ctxt) ths)); | |
| 648 | ||
| 649 | ||
| 650 | ||
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 651 | (** Hindley-Milner polymorphism **) | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 652 | |
| 7925 | 653 | (* warn_extra_tfrees *) | 
| 654 | ||
| 12130 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 655 | fun warn_extra_tfrees | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 656 |     (ctxt1 as Context {defs = (_, _, (names1, tab1)), ...})
 | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 657 |     (ctxt2 as Context {defs = (_, _, (names2, tab2)), ...}) =
 | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 658 | let | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 659 | fun known_tfree a (Type (_, Ts)) = exists (known_tfree a) Ts | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 660 | | known_tfree a (TFree (a', _)) = a = a' | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 661 | | known_tfree _ _ = false; | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 662 | |
| 12130 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 663 | val extras = | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 664 | Library.gen_rems Library.eq_fst (Symtab.dest tab2, Symtab.dest tab1) | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 665 | |> map (fn (a, ts) => map (pair a) (mapfilter (try (#1 o Term.dest_Free)) ts)) |> flat | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 666 | |> mapfilter (fn (a, x) => | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 667 | (case def_type ctxt1 false (x, ~1) of None => Some (a, x) | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 668 | | Some T => if known_tfree a T then None else Some (a, x))); | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 669 | val tfrees = map #1 extras |> Library.sort_strings |> Library.unique_strings; | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 670 | val frees = map #2 extras |> Library.sort_strings |> Library.unique_strings; | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 671 | in | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 672 | if null extras then () | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 673 |     else warning ("Just introduced free type variable(s): " ^ commas tfrees ^ " in " ^
 | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 674 | space_implode " or " frees); | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 675 | ctxt2 | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 676 | end; | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 677 | |
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 678 | |
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 679 | (* generalize type variables *) | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 680 | |
| 12008 | 681 | fun gen_tfrees inner outer = | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 682 | let | 
| 12057 | 683 | val extra_fixes = fixed_names_of inner \\ fixed_names_of outer; | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 684 | fun still_fixed (Free (x, _)) = not (x mem_string extra_fixes) | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 685 | | still_fixed _ = false; | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 686 | fun add (gen, (a, xs)) = | 
| 12008 | 687 | if is_some (Symtab.lookup (used_table outer, a)) orelse exists still_fixed xs | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 688 | then gen else a :: gen; | 
| 12008 | 689 | in Symtab.foldl add ([], used_table inner) end; | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 690 | |
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 691 | fun generalizeT inner outer = | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 692 | let | 
| 12008 | 693 | val tfrees = gen_tfrees inner outer; | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 694 | fun gen (x, S) = if x mem_string tfrees then TVar ((x, 0), S) else TFree (x, S); | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 695 | in Term.map_type_tfree gen end; | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 696 | |
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 697 | val generalize = Term.map_term_types oo generalizeT; | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 698 | |
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 699 | |
| 9553 | 700 | |
| 701 | (** export theorems **) | |
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 702 | |
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 703 | fun get_free x (None, t as Free (y, _)) = if x = y then Some t else None | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 704 | | get_free _ (opt, _) = opt; | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 705 | |
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 706 | fun find_free t x = foldl_aterms (get_free x) (None, t); | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 707 | |
| 12008 | 708 | fun export is_goal inner outer = | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 709 | let | 
| 12008 | 710 | val tfrees = gen_tfrees inner outer; | 
| 12057 | 711 | val fixes = fixed_names_of inner \\ fixed_names_of outer; | 
| 712 | val asms = Library.drop (length (assumptions_of outer), assumptions_of inner); | |
| 11816 | 713 | val exp_asms = map (fn (cprops, exp) => exp is_goal cprops) asms; | 
| 714 | in fn thm => | |
| 715 | thm | |
| 716 | |> Tactic.norm_hhf | |
| 717 | |> Seq.EVERY (rev exp_asms) | |
| 718 | |> Seq.map (fn rule => | |
| 719 | let | |
| 720 |         val {sign, prop, ...} = Thm.rep_thm rule;
 | |
| 721 | val frees = map (Thm.cterm_of sign) (mapfilter (find_free prop) fixes); | |
| 722 | in | |
| 723 | rule | |
| 724 | |> Drule.forall_intr_list frees | |
| 725 | |> Tactic.norm_hhf | |
| 726 | |> Drule.tvars_intr_list tfrees | |
| 727 | end) | |
| 728 | end; | |
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 729 | |
| 7925 | 730 | |
| 5819 | 731 | |
| 732 | (** bindings **) | |
| 733 | ||
| 734 | (* update_binds *) | |
| 735 | ||
| 7606 | 736 | fun del_bind (ctxt, xi) = | 
| 737 | ctxt | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 738 | |> map_context (fn (thy, syntax, data, asms, binds, thms, cases, defs) => | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 739 | (thy, syntax, data, asms, Vartab.update ((xi, None), binds), thms, cases, defs)); | 
| 7606 | 740 | |
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 741 | fun upd_bind (ctxt, ((x, i), t)) = | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 742 | let | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 743 | val T = Term.fastype_of t; | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 744 | val t' = | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 745 | if null (Term.term_tvars t \\ Term.typ_tvars T) then t | 
| 8637 | 746 | else Var ((x ^ "_has_extra_type_vars_on_rhs", i), T); | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 747 | in | 
| 5819 | 748 | ctxt | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 749 | |> declare_term t' | 
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 750 | |> map_context (fn (thy, syntax, data, asms, binds, thms, cases, defs) => | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 751 | (thy, syntax, data, asms, Vartab.update (((x, i), Some (t', T)), binds), thms, cases, defs)) | 
| 5819 | 752 | end; | 
| 753 | ||
| 7606 | 754 | fun del_upd_bind (ctxt, (xi, None)) = del_bind (ctxt, xi) | 
| 755 | | del_upd_bind (ctxt, (xi, Some t)) = upd_bind (ctxt, (xi, t)); | |
| 756 | ||
| 5819 | 757 | fun update_binds bs ctxt = foldl upd_bind (ctxt, bs); | 
| 7606 | 758 | fun delete_update_binds bs ctxt = foldl del_upd_bind (ctxt, bs); | 
| 759 | ||
| 5819 | 760 | |
| 8096 | 761 | (* simult_matches *) | 
| 762 | ||
| 10554 | 763 | local | 
| 764 | ||
| 765 | val add_vars = Term.foldl_aterms (fn (vs, Var (xi, _)) => xi ins vs | (vs, _) => vs); | |
| 766 | fun vars_of ts = foldl add_vars ([], ts); | |
| 767 | ||
| 768 | in | |
| 769 | ||
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 770 | fun simult_matches ctxt [] = [] | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 771 | | simult_matches ctxt pairs = | 
| 8096 | 772 | let | 
| 10554 | 773 |         fun fail () = raise CONTEXT ("Pattern match failed!", ctxt);
 | 
| 774 | ||
| 8096 | 775 | val maxidx = foldl (fn (i, (t1, t2)) => | 
| 776 | Int.max (i, Int.max (Term.maxidx_of_term t1, Term.maxidx_of_term t2))) (~1, pairs); | |
| 10554 | 777 | val envs = Unify.smash_unifiers (sign_of ctxt, Envir.empty maxidx, | 
| 778 | map swap pairs); (*prefer assignment of variables from patterns*) | |
| 779 | val env = | |
| 8096 | 780 | (case Seq.pull envs of | 
| 10554 | 781 | None => fail () | 
| 782 | | Some (env, _) => env); (*ignore further results*) | |
| 783 | val domain = filter_out Term.is_replaced_dummy_pattern (vars_of (map #1 pairs)); | |
| 784 | val _ = (*may not assign variables from text*) | |
| 785 | if null (map #1 (Envir.alist_of env) inter vars_of (map #2 pairs)) then () | |
| 786 | else fail (); | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 787 | fun norm_bind (xi, t) = if xi mem domain then Some (xi, Envir.norm_term env t) else None; | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 788 | in mapfilter norm_bind (Envir.alist_of env) end; | 
| 8096 | 789 | |
| 10554 | 790 | end; | 
| 791 | ||
| 8096 | 792 | |
| 793 | (* add_binds(_i) *) | |
| 5819 | 794 | |
| 7925 | 795 | local | 
| 796 | ||
| 5819 | 797 | fun gen_bind prep (ctxt, (xi as (x, _), raw_t)) = | 
| 7606 | 798 | ctxt |> delete_update_binds [(xi, apsome (prep ctxt) raw_t)]; | 
| 5819 | 799 | |
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 800 | fun gen_binds prep binds ctxt = foldl (gen_bind prep) (ctxt, binds); | 
| 7925 | 801 | |
| 10810 | 802 | in | 
| 803 | ||
| 10554 | 804 | fun drop_schematic (b as (xi, Some t)) = if null (Term.term_vars t) then b else (xi, None) | 
| 805 | | drop_schematic b = b; | |
| 806 | ||
| 5819 | 807 | val add_binds = gen_binds read_term; | 
| 808 | val add_binds_i = gen_binds cert_term; | |
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 809 | |
| 11764 | 810 | fun auto_bind_goal t ctxt = | 
| 811 | ctxt |> add_binds_i (map drop_schematic (AutoBind.goal (sign_of ctxt) t)); | |
| 812 | ||
| 813 | fun auto_bind_facts name ts ctxt = | |
| 814 | ctxt |> add_binds_i (map drop_schematic (AutoBind.facts (sign_of ctxt) name ts)); | |
| 7925 | 815 | |
| 816 | end; | |
| 5819 | 817 | |
| 818 | ||
| 8096 | 819 | (* match_bind(_i) *) | 
| 5819 | 820 | |
| 8096 | 821 | local | 
| 822 | ||
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 823 | fun prep_bind prep_pats (ctxt, (raw_pats, t)) = | 
| 5819 | 824 | let | 
| 8096 | 825 | val ctxt' = declare_term t ctxt; | 
| 826 | val pats = prep_pats (fastype_of t) ctxt' raw_pats; | |
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 827 | val binds = simult_matches ctxt' (map (rpair t) pats); | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 828 | in (ctxt', binds) end; | 
| 7670 | 829 | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 830 | fun gen_binds prep_terms prep_pats gen raw_binds ctxt = | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 831 | let | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 832 | val ts = prep_terms ctxt (map snd raw_binds); | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 833 | val (ctxt', binds) = | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 834 | apsnd flat (foldl_map (prep_bind prep_pats) (ctxt, map fst raw_binds ~~ ts)); | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 835 | val binds' = | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 836 | if gen then map (apsnd (generalize ctxt' ctxt)) binds | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 837 | else binds; | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 838 | val binds'' = map (apsnd Some) binds'; | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 839 | in | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 840 | warn_extra_tfrees ctxt | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 841 | (if gen then ctxt (*sic!*) |> declare_terms (map #2 binds') |> add_binds_i binds'' | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 842 | else ctxt' |> add_binds_i binds'') | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 843 | end; | 
| 8096 | 844 | |
| 845 | in | |
| 5935 | 846 | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 847 | val match_bind = gen_binds read_terms read_term_pats; | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 848 | val match_bind_i = gen_binds (map o cert_term) cert_term_pats; | 
| 8096 | 849 | |
| 850 | end; | |
| 5935 | 851 | |
| 852 | ||
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 853 | (* propositions with patterns *) | 
| 5935 | 854 | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 855 | local | 
| 8096 | 856 | |
| 10554 | 857 | fun prep_propp schematic prep_props prep_pats (context, args) = | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 858 | let | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 859 | fun prep ((ctxt, prop :: props), (_, (raw_pats1, raw_pats2))) = | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 860 | let | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 861 | val ctxt' = declare_term prop ctxt; | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 862 | val pats = prep_pats ctxt' (raw_pats1 @ raw_pats2); (*simultaneous type inference!*) | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 863 | val len1 = length raw_pats1; | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 864 | in ((ctxt', props), (prop, (take (len1, pats), drop (len1, pats)))) end | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 865 | | prep _ = sys_error "prep_propp"; | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 866 | val ((context', _), propp) = foldl_map (foldl_map prep) | 
| 10554 | 867 | ((context, prep_props schematic context (flat (map (map fst) args))), args); | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 868 | in (context', propp) end; | 
| 5935 | 869 | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 870 | fun matches ctxt (prop, (pats1, pats2)) = | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 871 | simult_matches ctxt (map (rpair prop) pats1 @ map (rpair (Logic.strip_imp_concl prop)) pats2); | 
| 8096 | 872 | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 873 | fun gen_bind_propp prepp (ctxt, raw_args) = | 
| 8096 | 874 | let | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 875 | val (ctxt', args) = prepp (ctxt, raw_args); | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 876 | val binds = flat (flat (map (map (matches ctxt')) args)); | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 877 | val propss = map (map #1) args; | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 878 | |
| 10554 | 879 | (*generalize result: context evaluated now, binds added later*) | 
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 880 | val gen = generalize ctxt' ctxt; | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 881 | fun gen_binds c = c |> add_binds_i (map (apsnd (Some o gen)) binds); | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 882 | in (ctxt' |> add_binds_i (map (apsnd Some) binds), (propss, gen_binds)) end; | 
| 8096 | 883 | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 884 | in | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 885 | |
| 11925 | 886 | val read_propp = prep_propp false read_props read_prop_pats; | 
| 887 | val cert_propp = prep_propp false cert_props cert_prop_pats; | |
| 10554 | 888 | val read_propp_schematic = prep_propp true read_props read_prop_pats; | 
| 889 | val cert_propp_schematic = prep_propp true cert_props cert_prop_pats; | |
| 890 | ||
| 11925 | 891 | val bind_propp = gen_bind_propp read_propp; | 
| 892 | val bind_propp_i = gen_bind_propp cert_propp; | |
| 893 | val bind_propp_schematic = gen_bind_propp read_propp_schematic; | |
| 10554 | 894 | val bind_propp_schematic_i = gen_bind_propp cert_propp_schematic; | 
| 6789 | 895 | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 896 | end; | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 897 | |
| 6789 | 898 | |
| 5819 | 899 | |
| 900 | (** theorems **) | |
| 901 | ||
| 6091 | 902 | (* get_thm(s) *) | 
| 5819 | 903 | |
| 9566 | 904 | (*beware of proper order of evaluation!*) | 
| 5819 | 905 | |
| 9566 | 906 | fun retrieve_thms f g (ctxt as Context {thy, thms, ...}) =
 | 
| 907 | let | |
| 908 | val sg_ref = Sign.self_ref (Theory.sign_of thy); | |
| 909 | val get_from_thy = f thy; | |
| 910 | in | |
| 911 | fn name => | |
| 912 | (case Symtab.lookup (thms, name) of | |
| 913 | Some (Some ths) => map (Thm.transfer_sg (Sign.deref sg_ref)) ths | |
| 914 | | _ => get_from_thy name) |> g name | |
| 915 | end; | |
| 5819 | 916 | |
| 9566 | 917 | val get_thm = retrieve_thms PureThy.get_thms PureThy.single_thm; | 
| 918 | val get_thm_closure = retrieve_thms PureThy.get_thms_closure PureThy.single_thm; | |
| 919 | val get_thms = retrieve_thms PureThy.get_thms (K I); | |
| 920 | val get_thms_closure = retrieve_thms PureThy.get_thms_closure (K I); | |
| 5819 | 921 | |
| 922 | ||
| 6091 | 923 | (* put_thm(s) *) | 
| 5819 | 924 | |
| 7479 | 925 | fun put_thms ("", _) = I
 | 
| 926 | | put_thms (name, ths) = map_context | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 927 | (fn (thy, syntax, data, asms, binds, thms, cases, defs) => | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 928 | (thy, syntax, data, asms, binds, Symtab.update ((name, Some ths), thms), cases, defs)); | 
| 5819 | 929 | |
| 6091 | 930 | fun put_thm (name, th) = put_thms (name, [th]); | 
| 5819 | 931 | |
| 6091 | 932 | fun put_thmss [] ctxt = ctxt | 
| 933 | | put_thmss (th :: ths) ctxt = ctxt |> put_thms th |> put_thmss ths; | |
| 5819 | 934 | |
| 935 | ||
| 7606 | 936 | (* reset_thms *) | 
| 937 | ||
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 938 | fun reset_thms name = map_context (fn (thy, syntax, data, asms, binds, thms, cases, defs) => | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 939 | (thy, syntax, data, asms, binds, Symtab.update ((name, None), thms), cases, defs)); | 
| 7606 | 940 | |
| 941 | ||
| 6091 | 942 | (* have_thmss *) | 
| 5819 | 943 | |
| 9196 | 944 | fun have_thss (ctxt, ((name, more_attrs), ths_attrs)) = | 
| 5819 | 945 | let | 
| 946 | fun app ((ct, ths), (th, attrs)) = | |
| 6091 | 947 | let val (ct', th') = Thm.applys_attributes ((ct, th), attrs @ more_attrs) | 
| 5819 | 948 | in (ct', th' :: ths) end | 
| 949 | val (ctxt', rev_thms) = foldl app ((ctxt, []), ths_attrs); | |
| 9196 | 950 | val thms = flat (rev rev_thms); | 
| 6091 | 951 | in (ctxt' |> put_thms (name, thms), (name, thms)) end; | 
| 5819 | 952 | |
| 9196 | 953 | fun have_thmss args ctxt = foldl_map have_thss (ctxt, args); | 
| 954 | ||
| 5819 | 955 | |
| 956 | ||
| 957 | (** assumptions **) | |
| 958 | ||
| 11918 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 959 | (* basic exporters *) | 
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 960 | |
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 961 | fun export_assume true = Seq.single oo Drule.implies_intr_goals | 
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 962 | | export_assume false = Seq.single oo Drule.implies_intr_list; | 
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 963 | |
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 964 | fun export_presume _ = export_assume false; | 
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 965 | |
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 966 | |
| 12057 | 967 | (* defs *) | 
| 968 | ||
| 12066 | 969 | fun cert_def ctxt eq = | 
| 12016 | 970 | let | 
| 12057 | 971 | fun err msg = raise CONTEXT (msg ^ | 
| 972 | "\nThe error(s) above occurred in local definition: " ^ string_of_term ctxt eq, ctxt); | |
| 973 | val (lhs, rhs) = Logic.dest_equals (Term.strip_all_body eq) | |
| 974 | handle TERM _ => err "Not a meta-equality (==)"; | |
| 12086 | 975 | val (f, xs) = Term.strip_comb lhs; | 
| 976 | val (c, _) = Term.dest_Free f handle TERM _ => | |
| 977 | err "Head of lhs must be free or fixed variable"; | |
| 12057 | 978 | |
| 12086 | 979 | fun is_free (Free (x, _)) = not (is_fixed ctxt x) | 
| 980 | | is_free _ = false; | |
| 981 | val extra_frees = filter is_free (term_frees rhs) \\ xs; | |
| 12016 | 982 | in | 
| 12086 | 983 | conditional (not (forall (is_Bound orf is_free) xs andalso null (duplicates xs))) (fn () => | 
| 984 | err "Arguments of lhs must be distinct free or fixed variables"); | |
| 985 | conditional (f mem Term.term_frees rhs) (fn () => | |
| 986 | err "Element to be defined occurs on rhs"); | |
| 987 | conditional (not (null extra_frees)) (fn () => | |
| 988 |       err ("Extra free variables on rhs: " ^ commas_quote (map (#1 o dest_Free) extra_frees)));
 | |
| 989 | (c, Term.list_all_free (mapfilter (try Term.dest_Free) xs, eq)) | |
| 12057 | 990 | end; | 
| 991 | ||
| 992 | fun head_of_def cprop = | |
| 993 | #1 (Term.strip_comb (#1 (Logic.dest_equals (Term.strip_all_body (Thm.term_of cprop))))) | |
| 994 | |> Thm.cterm_of (Thm.sign_of_cterm cprop); | |
| 11918 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 995 | |
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 996 | fun export_def _ cprops thm = | 
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 997 | thm | 
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 998 | |> Drule.implies_intr_list cprops | 
| 12057 | 999 | |> Drule.forall_intr_list (map head_of_def cprops) | 
| 11918 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 1000 | |> Drule.forall_elim_vars 0 | 
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 1001 | |> RANGE (replicate (length cprops) (Tactic.rtac Drule.reflexive_thm)) 1; | 
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 1002 | |
| 
dfdf0798d7b8
added export_assume, export_presume, export_def (from proof.ML);
 wenzelm parents: 
11915diff
changeset | 1003 | |
| 5819 | 1004 | (* assume *) | 
| 1005 | ||
| 7270 | 1006 | local | 
| 6797 | 1007 | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 1008 | fun add_assm (ctxt, ((name, attrs), props)) = | 
| 5819 | 1009 | let | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 1010 | val cprops = map (Thm.cterm_of (sign_of ctxt)) props; | 
| 11816 | 1011 | val asms = map (Tactic.norm_hhf o Thm.assume) cprops; | 
| 5919 | 1012 | |
| 1013 | val ths = map (fn th => ([th], [])) asms; | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 1014 | val (ctxt', [(_, thms)]) = | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 1015 | ctxt | 
| 6797 | 1016 | |> auto_bind_facts name props | 
| 11526 | 1017 | |> have_thmss [((name, attrs), ths)]; | 
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 1018 | in (ctxt', (cprops, (name, asms), (name, thms))) end; | 
| 5819 | 1019 | |
| 9470 | 1020 | fun gen_assms prepp exp args ctxt = | 
| 1021 | let | |
| 10465 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 1022 | val (ctxt1, propss) = prepp (ctxt, map snd args); | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 1023 | val (ctxt2, results) = foldl_map add_assm (ctxt1, map fst args ~~ propss); | 
| 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 wenzelm parents: 
10381diff
changeset | 1024 | |
| 9470 | 1025 | val cprops = flat (map #1 results); | 
| 1026 | val asmss = map #2 results; | |
| 1027 | val thmss = map #3 results; | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1028 | val ctxt3 = ctxt2 |> map_context | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1029 | (fn (thy, syntax, data, ((asms_ct, asms_th), fixes), binds, thms, cases, defs) => | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1030 | (thy, syntax, data, ((asms_ct @ [(cprops, exp)], asms_th @ asmss), fixes), binds, thms, | 
| 9470 | 1031 | cases, defs)); | 
| 11925 | 1032 |     val ctxt4 = ctxt3 |> put_thms ("prems", prems_of ctxt3);
 | 
| 1033 | in (warn_extra_tfrees ctxt ctxt4, thmss) end; | |
| 5819 | 1034 | |
| 7270 | 1035 | in | 
| 1036 | ||
| 8616 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 1037 | val assume = gen_assms (apsnd #1 o bind_propp); | 
| 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8462diff
changeset | 1038 | val assume_i = gen_assms (apsnd #1 o bind_propp_i); | 
| 7270 | 1039 | |
| 1040 | end; | |
| 5819 | 1041 | |
| 1042 | ||
| 8096 | 1043 | (* variables *) | 
| 1044 | ||
| 10381 | 1045 | local | 
| 1046 | ||
| 11925 | 1047 | fun prep_vars prep_typ no_internal (ctxt, (xs, raw_T)) = | 
| 8096 | 1048 | let | 
| 11925 | 1049 | val _ = (case filter (not o Syntax.is_identifier) (map (no_skolem no_internal ctxt) xs) of | 
| 8096 | 1050 |       [] => () | bads => raise CONTEXT ("Bad variable name(s): " ^ commas_quote bads, ctxt));
 | 
| 1051 | ||
| 1052 | val opt_T = apsome (prep_typ ctxt) raw_T; | |
| 1053 | val T = if_none opt_T TypeInfer.logicT; | |
| 12130 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1054 | val ctxt' = ctxt |> declare_terms_syntax (map (fn x => Free (x, T)) xs); | 
| 8096 | 1055 | in (ctxt', (xs, opt_T)) end; | 
| 1056 | ||
| 10381 | 1057 | in | 
| 1058 | ||
| 11925 | 1059 | val read_vars = prep_vars read_typ true; | 
| 1060 | val cert_vars = prep_vars cert_typ false; | |
| 8096 | 1061 | |
| 10381 | 1062 | end; | 
| 1063 | ||
| 8096 | 1064 | |
| 5819 | 1065 | (* fix *) | 
| 1066 | ||
| 8096 | 1067 | local | 
| 1068 | ||
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1069 | fun map_vars f = map_context (fn (thy, syntax, data, (assumes, vars), binds, thms, cases, defs) => | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1070 | (thy, syntax, data, (assumes, f vars), binds, thms, cases, defs)); | 
| 5819 | 1071 | |
| 11925 | 1072 | fun err_dups ctxt xs = raise CONTEXT ("Duplicate variable(s): " ^ commas_quote xs, ctxt);
 | 
| 1073 | ||
| 12130 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1074 | val declare = | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1075 | declare_terms_syntax o mapfilter (fn (_, None) => None | (x, Some T) => Some (Free (x, T))); | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1076 | |
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1077 | fun add_vars xs Ts ctxt = | 
| 11925 | 1078 | let | 
| 12130 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1079 | val xs' = variantlist (xs, #2 (vars_of ctxt)); | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1080 | val xs'' = map Syntax.skolem xs'; | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1081 | in | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1082 | ctxt |> declare (xs'' ~~ Ts) | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1083 | |> map_vars (fn (fixes, used) => ((xs ~~ xs'') @ fixes, xs' @ used)) | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1084 | end; | 
| 11925 | 1085 | |
| 12130 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1086 | fun add_vars_direct xs Ts ctxt = | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1087 | ctxt |> declare (xs ~~ Ts) | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1088 | |> map_vars (fn (fixes, used) => | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1089 | (case xs inter_string map #1 fixes of | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1090 | [] => ((xs ~~ xs) @ fixes, xs @ used) | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1091 | | dups => err_dups ctxt dups)); | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1092 | |
| 11925 | 1093 | |
| 1094 | fun gen_fix prep add raw_vars ctxt = | |
| 8096 | 1095 | let | 
| 1096 | val (ctxt', varss) = foldl_map prep (ctxt, raw_vars); | |
| 12130 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1097 | val vars = rev (flat (map (fn (xs, T) => map (rpair T) xs) varss)); | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1098 | val xs = map #1 vars; | 
| 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1099 | val Ts = map #2 vars; | 
| 8096 | 1100 | in | 
| 11925 | 1101 | (case Library.duplicates xs of [] => () | dups => err_dups ctxt dups); | 
| 12130 
30d9143aff7e
syntactic declaration of external *and* internal versions of fixes;
 wenzelm parents: 
12123diff
changeset | 1102 | ctxt' |> add xs Ts | 
| 8096 | 1103 | end; | 
| 5819 | 1104 | |
| 8096 | 1105 | in | 
| 7679 | 1106 | |
| 11925 | 1107 | val fix = gen_fix read_vars add_vars; | 
| 1108 | val fix_i = gen_fix cert_vars add_vars; | |
| 1109 | val fix_direct = gen_fix cert_vars add_vars_direct; | |
| 8096 | 1110 | |
| 1111 | end; | |
| 5819 | 1112 | |
| 12048 | 1113 | fun fix_frees ts ctxt = | 
| 1114 | let | |
| 12057 | 1115 | val frees = foldl Drule.add_frees ([], ts); | 
| 1116 | fun new (x, T) = if is_fixed ctxt x then None else Some ([x], Some T); | |
| 12048 | 1117 | in fix_direct (mapfilter new frees) ctxt end; | 
| 12016 | 1118 | |
| 6895 | 1119 | |
| 9291 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1120 | (*Note: improper use may result in variable capture / dynamic scoping!*) | 
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1121 | fun bind_skolem ctxt xs = | 
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1122 | let | 
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1123 | val ctxt' = ctxt |> fix_i [(xs, None)]; | 
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1124 | fun bind (t as Free (x, T)) = | 
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1125 | if x mem_string xs then | 
| 10583 | 1126 | (case lookup_skolem ctxt' x of Some x' => Free (x', T) | None => t) | 
| 9291 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1127 | else t | 
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1128 | | bind (t $ u) = bind t $ bind u | 
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1129 | | bind (Abs (x, T, t)) = Abs (x, T, bind t) | 
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1130 | | bind a = a; | 
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1131 | in bind end; | 
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1132 | |
| 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 wenzelm parents: 
9274diff
changeset | 1133 | |
| 5819 | 1134 | |
| 8373 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 wenzelm parents: 
8186diff
changeset | 1135 | (** cases **) | 
| 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 wenzelm parents: 
8186diff
changeset | 1136 | |
| 11793 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1137 | fun prep_case ctxt name xs {fixes, assumes, binds} =
 | 
| 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1138 | let | 
| 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1139 | fun replace (opt_x :: xs) ((y, T) :: ys) = (if_none opt_x y, T) :: replace xs ys | 
| 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1140 | | replace [] ys = ys | 
| 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1141 |       | replace (_ :: _) [] = raise CONTEXT ("Too many parameters for case " ^ quote name, ctxt);
 | 
| 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1142 | in | 
| 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1143 | if null (foldr Term.add_typ_tvars (map snd fixes, [])) andalso | 
| 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1144 | null (foldr Term.add_term_vars (assumes, [])) then | 
| 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1145 |         {fixes = replace xs fixes, assumes = assumes, binds = map drop_schematic binds}
 | 
| 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1146 |     else raise CONTEXT ("Illegal schematic variable(s) in case " ^ quote name, ctxt)
 | 
| 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1147 | end; | 
| 8403 | 1148 | |
| 11793 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1149 | fun get_case (ctxt as Context {cases, ...}) name xs =
 | 
| 8426 | 1150 | (case assoc (cases, name) of | 
| 8373 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 wenzelm parents: 
8186diff
changeset | 1151 |     None => raise CONTEXT ("Unknown case: " ^ quote name, ctxt)
 | 
| 11793 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11764diff
changeset | 1152 | | Some c => prep_case ctxt name xs c); | 
| 8373 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 wenzelm parents: 
8186diff
changeset | 1153 | |
| 8384 | 1154 | |
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1155 | fun add_cases xs = map_context (fn (thy, syntax, data, asms, binds, thms, cases, defs) => | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1156 | (thy, syntax, data, asms, binds, thms, rev (filter_out (equal "" o #1) xs) @ cases, defs)); | 
| 8373 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 wenzelm parents: 
8186diff
changeset | 1157 | |
| 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 wenzelm parents: 
8186diff
changeset | 1158 | |
| 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 wenzelm parents: 
8186diff
changeset | 1159 | |
| 10810 | 1160 | (** print context information **) | 
| 1161 | ||
| 1162 | val verbose = ref false; | |
| 1163 | fun verb f x = if ! verbose then f (x ()) else []; | |
| 1164 | fun verb_single x = verb Library.single x; | |
| 1165 | ||
| 1166 | fun setmp_verbose f x = Library.setmp verbose true f x; | |
| 1167 | ||
| 1168 | fun pretty_items prt name items = | |
| 1169 | let | |
| 1170 | fun prt_itms (name, [x]) = Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, prt x] | |
| 1171 | | prt_itms (name, xs) = Pretty.big_list (name ^ ":") (map prt xs); | |
| 1172 | in | |
| 1173 | if null items andalso not (! verbose) then [] | |
| 1174 | else [Pretty.big_list name (map prt_itms items)] | |
| 1175 | end; | |
| 1176 | ||
| 1177 | ||
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1178 | (* local syntax *) | 
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1179 | |
| 12093 | 1180 | val print_syntax = Syntax.print_syntax o syn_of; | 
| 12072 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1181 | |
| 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 wenzelm parents: 
12066diff
changeset | 1182 | |
| 10810 | 1183 | (* term bindings *) | 
| 1184 | ||
| 1185 | val smash_option = fn (_, None) => None | (xi, Some b) => Some (xi, b); | |
| 1186 | ||
| 1187 | fun pretty_binds (ctxt as Context {binds, ...}) =
 | |
| 1188 | let | |
| 12057 | 1189 | fun prt_bind (xi, (t, T)) = pretty_term ctxt (Logic.mk_equals (Var (xi, T), t)); | 
| 10810 | 1190 | val bs = mapfilter smash_option (Vartab.dest binds); | 
| 1191 | in | |
| 1192 | if null bs andalso not (! verbose) then [] | |
| 1193 | else [Pretty.big_list "term bindings:" (map prt_bind bs)] | |
| 1194 | end; | |
| 1195 | ||
| 1196 | val print_binds = Pretty.writeln o Pretty.chunks o pretty_binds; | |
| 1197 | ||
| 1198 | ||
| 1199 | (* local theorems *) | |
| 1200 | ||
| 12057 | 1201 | fun pretty_lthms (ctxt as Context {thms, ...}) =
 | 
| 1202 | pretty_items (pretty_thm ctxt) "facts:" (mapfilter smash_option (Symtab.dest thms)); | |
| 10810 | 1203 | |
| 12057 | 1204 | val print_lthms = Pretty.writeln o Pretty.chunks o pretty_lthms; | 
| 10810 | 1205 | |
| 1206 | ||
| 1207 | (* local contexts *) | |
| 1208 | ||
| 10830 | 1209 | fun apply_case ({fixes, assumes, binds}: RuleCases.T) ctxt =
 | 
| 10810 | 1210 | let | 
| 10830 | 1211 | fun bind (c, (x, T)) = (c |> fix_i [([x], Some T)], bind_skolem c [x] (Free (x, T))); | 
| 1212 | val (ctxt', xs) = foldl_map bind (ctxt, fixes); | |
| 10810 | 1213 | fun app t = foldl Term.betapply (t, xs); | 
| 10830 | 1214 | in (ctxt', (map (apsnd (apsome app)) binds, map app assumes)) end; | 
| 10810 | 1215 | |
| 1216 | fun pretty_cases (ctxt as Context {cases, ...}) =
 | |
| 1217 | let | |
| 12057 | 1218 | val prt_term = pretty_term ctxt; | 
| 1219 | ||
| 10810 | 1220 | fun prt_let (xi, t) = Pretty.block | 
| 10818 | 1221 | [Pretty.quote (prt_term (Var (xi, Term.fastype_of t))), Pretty.str " =", Pretty.brk 1, | 
| 10810 | 1222 | Pretty.quote (prt_term t)]; | 
| 1223 | ||
| 1224 | fun prt_sect _ _ _ [] = [] | |
| 1225 | | prt_sect s sep prt xs = [Pretty.block (Pretty.breaks (Pretty.str s :: | |
| 1226 | flat (Library.separate sep (map (Library.single o prt) xs))))]; | |
| 1227 | ||
| 10830 | 1228 | fun prt_case (name, (fixes, (lets, asms))) = Pretty.block (Pretty.fbreaks | 
| 10810 | 1229 | (Pretty.str (name ^ ":") :: | 
| 11915 | 1230 | prt_sect "fix" [] (Pretty.str o fst) fixes @ | 
| 10810 | 1231 | prt_sect "let" [Pretty.str "and"] prt_let | 
| 1232 | (mapfilter (fn (xi, Some t) => Some (xi, t) | _ => None) lets) @ | |
| 1233 | prt_sect "assume" [] (Pretty.quote o prt_term) asms)); | |
| 1234 | ||
| 1235 | val cases' = rev (Library.gen_distinct Library.eq_fst cases); | |
| 1236 | in | |
| 1237 | if null cases andalso not (! verbose) then [] | |
| 10830 | 1238 | else [Pretty.big_list "cases:" | 
| 1239 | (map (prt_case o apsnd (fn c => (#fixes c, #2 (apply_case c ctxt)))) cases')] | |
| 10810 | 1240 | end; | 
| 1241 | ||
| 1242 | val print_cases = Pretty.writeln o Pretty.chunks o pretty_cases; | |
| 1243 | ||
| 1244 | ||
| 12057 | 1245 | (* core context *) | 
| 10810 | 1246 | |
| 1247 | val prems_limit = ref 10; | |
| 1248 | ||
| 12057 | 1249 | fun pretty_asms ctxt = | 
| 10810 | 1250 | let | 
| 12057 | 1251 | val prt_term = pretty_term ctxt; | 
| 1252 | ||
| 12093 | 1253 | (*structures*) | 
| 1254 | val (_, structs, _) = syntax_of ctxt; | |
| 1255 | val prt_structs = if null structs then [] | |
| 1256 | else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 :: | |
| 1257 | Pretty.commas (map Pretty.str structs))]; | |
| 1258 | ||
| 12057 | 1259 | (*fixes*) | 
| 1260 | fun prt_fix (x, x') = | |
| 1261 | if x = x' then Pretty.str x | |
| 1262 | else Pretty.block [Pretty.str x, Pretty.str " =", Pretty.brk 1, prt_term (Syntax.free x')]; | |
| 12093 | 1263 | val fixes = rev (filter_out | 
| 1264 | ((can Syntax.dest_internal o #1) orf (fn (_, x') => x' mem_string structs)) (fixes_of ctxt)); | |
| 1265 | val prt_fixes = if null fixes then [] | |
| 1266 | else [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 :: | |
| 1267 | Pretty.commas (map prt_fix fixes))]; | |
| 12057 | 1268 | |
| 1269 | (*prems*) | |
| 10810 | 1270 | val limit = ! prems_limit; | 
| 1271 | val prems = prems_of ctxt; | |
| 1272 | val len = length prems; | |
| 12093 | 1273 | val prt_prems = if null prems then [] | 
| 1274 | else [Pretty.big_list "prems:" ((if len <= limit then [] else [Pretty.str "..."]) @ | |
| 1275 | map (pretty_thm ctxt) (Library.drop (len - limit, prems)))]; | |
| 1276 | ||
| 1277 | in prt_structs @ prt_fixes @ prt_prems end; | |
| 10810 | 1278 | |
| 1279 | ||
| 1280 | (* main context *) | |
| 1281 | ||
| 12057 | 1282 | fun pretty_context (ctxt as Context {cases, defs = (types, sorts, (used, _)), ...}) =
 | 
| 10810 | 1283 | let | 
| 12057 | 1284 | val prt_term = pretty_term ctxt; | 
| 1285 | val prt_typ = pretty_typ ctxt; | |
| 1286 | val prt_sort = pretty_sort ctxt; | |
| 10810 | 1287 | |
| 1288 | (*theory*) | |
| 12057 | 1289 | val pretty_thy = Pretty.block | 
| 1290 | [Pretty.str "Theory:", Pretty.brk 1, Sign.pretty_sg (sign_of ctxt)]; | |
| 10810 | 1291 | |
| 1292 | (*defaults*) | |
| 1293 | fun prt_atom prt prtT (x, X) = Pretty.block | |
| 1294 | [prt x, Pretty.str " ::", Pretty.brk 1, prtT X]; | |
| 1295 | ||
| 1296 | fun prt_var (x, ~1) = prt_term (Syntax.free x) | |
| 1297 | | prt_var xi = prt_term (Syntax.var xi); | |
| 1298 | ||
| 1299 | fun prt_varT (x, ~1) = prt_typ (TFree (x, [])) | |
| 1300 | | prt_varT xi = prt_typ (TVar (xi, [])); | |
| 1301 | ||
| 1302 | val prt_defT = prt_atom prt_var prt_typ; | |
| 1303 | val prt_defS = prt_atom prt_varT prt_sort; | |
| 1304 | in | |
| 1305 | verb_single (K pretty_thy) @ | |
| 12057 | 1306 | pretty_asms ctxt @ | 
| 10810 | 1307 | verb pretty_binds (K ctxt) @ | 
| 12057 | 1308 | verb pretty_lthms (K ctxt) @ | 
| 10810 | 1309 | verb pretty_cases (K ctxt) @ | 
| 1310 | verb_single (fn () => Pretty.big_list "type constraints:" (map prt_defT (Vartab.dest types))) @ | |
| 1311 | verb_single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts))) @ | |
| 1312 |     verb_single (fn () => Pretty.strs ("used type variable names:" :: used))
 | |
| 1313 | end; | |
| 1314 | ||
| 1315 | ||
| 1316 | ||
| 5819 | 1317 | (** theory setup **) | 
| 1318 | ||
| 1319 | val setup = [ProofDataData.init]; | |
| 1320 | ||
| 1321 | end; |