| author | Fabian Huch <huch@in.tum.de> | 
| Wed, 05 Jun 2024 20:09:04 +0200 | |
| changeset 80257 | 96cb31f0bbdf | 
| parent 79336 | 032a31db4c6f | 
| child 80888 | c5ea0cb4dd91 | 
| permissions | -rw-r--r-- | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 1 | (* Title: Pure/Isar/expression.ML | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 2 | Author: Clemens Ballarin, TU Muenchen | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 3 | |
| 32800 | 4 | Locale expressions and user interface layer of locales. | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 5 | *) | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 6 | |
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 7 | signature EXPRESSION = | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 8 | sig | 
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 9 | (* Locale expressions *) | 
| 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 10 | datatype 'term map = Positional of 'term option list | Named of (string * 'term) list | 
| 67740 | 11 | type 'term rewrites = (Attrib.binding * 'term) list | 
| 12 |   type ('name, 'term) expr = ('name * ((string * bool) * ('term map * 'term rewrites))) list
 | |
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 13 | type expression_i = (string, term) expr * (binding * typ option * mixfix) list | 
| 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 14 | type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 15 | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 16 | (* Processing of context statements *) | 
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 17 | val cert_statement: Element.context_i list -> Element.statement_i -> | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 18 | Proof.context -> (Attrib.binding * (term * term list) list) list * Proof.context | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 19 | val read_statement: Element.context list -> Element.statement -> | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 20 | Proof.context -> (Attrib.binding * (term * term list) list) list * Proof.context | 
| 28879 | 21 | |
| 28795 | 22 | (* Declaring locales *) | 
| 55639 | 23 | val cert_declaration: expression_i -> (Proof.context -> Proof.context) -> | 
| 24 | Element.context_i list -> | |
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 25 | Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list | 
| 47311 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 wenzelm parents: 
47287diff
changeset | 26 | * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) | 
| 55639 | 27 | val cert_read_declaration: expression_i -> (Proof.context -> Proof.context) -> | 
| 28 | Element.context list -> | |
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 29 | Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list | 
| 47311 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 wenzelm parents: 
47287diff
changeset | 30 | * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) | 
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 31 | (*FIXME*) | 
| 29702 | 32 | val read_declaration: expression -> (Proof.context -> Proof.context) -> Element.context list -> | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 33 | Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list | 
| 47311 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 wenzelm parents: 
47287diff
changeset | 34 | * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) | 
| 72605 | 35 | val add_locale: binding -> binding -> Bundle.name list -> | 
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
41435diff
changeset | 36 | expression_i -> Element.context_i list -> theory -> string * local_theory | 
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 37 | val add_locale_cmd: binding -> binding -> (xstring * Position.T) list -> | 
| 41585 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 wenzelm parents: 
41435diff
changeset | 38 | expression -> Element.context list -> theory -> string * local_theory | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 39 | |
| 61669 | 40 | (* Processing of locale expressions *) | 
| 29441 | 41 | val cert_goal_expression: expression_i -> Proof.context -> | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 42 | (term list list * term list list * (string * morphism) list * (Attrib.binding * term) list list * morphism) * Proof.context | 
| 29496 | 43 | val read_goal_expression: expression -> Proof.context -> | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 44 | (term list list * term list list * (string * morphism) list * (Attrib.binding * term) list list * morphism) * Proof.context | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 45 | end; | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 46 | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 47 | structure Expression : EXPRESSION = | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 48 | struct | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 49 | |
| 28795 | 50 | datatype ctxt = datatype Element.ctxt; | 
| 51 | ||
| 52 | ||
| 53 | (*** Expressions ***) | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 54 | |
| 28872 | 55 | datatype 'term map = | 
| 56 | Positional of 'term option list | | |
| 57 | Named of (string * 'term) list; | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 58 | |
| 67740 | 59 | type 'term rewrites = (Attrib.binding * 'term) list; | 
| 60 | ||
| 61 | type ('name, 'term) expr = ('name * ((string * bool) * ('term map * 'term rewrites))) list;
 | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 62 | |
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 63 | type expression_i = (string, term) expr * (binding * typ option * mixfix) list; | 
| 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 64 | type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list; | 
| 28795 | 65 | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 66 | |
| 28859 | 67 | (** Internalise locale names in expr **) | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 68 | |
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 69 | fun check_expr thy instances = map (apfst (Locale.check thy)) instances; | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 70 | |
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 71 | |
| 30778 | 72 | (** Parameters of expression **) | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 73 | |
| 30778 | 74 | (*Sanity check of instantiations and extraction of implicit parameters. | 
| 75 | The latter only occurs iff strict = false. | |
| 76 | Positional instantiations are extended to match full length of parameter list | |
| 77 | of instantiated locale.*) | |
| 28895 | 78 | |
| 79 | fun parameters_of thy strict (expr, fixed) = | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 80 | let | 
| 53041 | 81 | val ctxt = Proof_Context.init_global thy; | 
| 82 | ||
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 83 | fun reject_dups message xs = | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 84 | (case duplicates (op =) xs of | 
| 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 85 | [] => () | 
| 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 86 | | dups => error (message ^ commas dups)); | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 87 | |
| 62752 | 88 | fun parm_eq ((p1, mx1), (p2, mx2)) = | 
| 89 | p1 = p2 andalso | |
| 90 | (Mixfix.equal (mx1, mx2) orelse | |
| 91 |           error ("Conflicting syntax for parameter " ^ quote p1 ^ " in expression" ^
 | |
| 92 | Position.here_list [Mixfix.pos_of mx1, Mixfix.pos_of mx2])); | |
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 93 | |
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 94 | fun params_loc loc = Locale.params_of thy loc |> map (apfst #1); | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 95 | fun params_inst (loc, (prfx, (Positional insts, eqns))) = | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 96 | let | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 97 | val ps = params_loc loc; | 
| 29358 | 98 | val d = length ps - length insts; | 
| 99 | val insts' = | |
| 53041 | 100 | if d < 0 then | 
| 101 |                 error ("More arguments than parameters in instantiation of locale " ^
 | |
| 102 | quote (Locale.markup_name ctxt loc)) | |
| 29358 | 103 | else insts @ replicate d NONE; | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 104 | val ps' = (ps ~~ insts') |> | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 105 | map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE); | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 106 | in (ps', (loc, (prfx, (Positional insts', eqns)))) end | 
| 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 107 | | params_inst (loc, (prfx, (Named insts, eqns))) = | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 108 | let | 
| 53041 | 109 | val _ = | 
| 110 | reject_dups "Duplicate instantiation of the following parameter(s): " | |
| 111 | (map fst insts); | |
| 30778 | 112 | val ps' = (insts, params_loc loc) |-> fold (fn (p, _) => fn ps => | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 113 | if AList.defined (op =) ps p then AList.delete (op =) p ps | 
| 30778 | 114 | else error (quote p ^ " not a parameter of instantiated expression")); | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 115 | in (ps', (loc, (prfx, (Named insts, eqns)))) end; | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 116 | fun params_expr is = | 
| 30778 | 117 | let | 
| 118 | val (is', ps') = fold_map (fn i => fn ps => | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 119 | let | 
| 30778 | 120 | val (ps', i') = params_inst i; | 
| 121 | val ps'' = distinct parm_eq (ps @ ps'); | |
| 122 | in (i', ps'') end) is [] | |
| 123 | in (ps', is') end; | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 124 | |
| 28895 | 125 | val (implicit, expr') = params_expr expr; | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 126 | |
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 127 | val implicit' = map #1 implicit; | 
| 42494 | 128 | val fixed' = map (Variable.check_name o #1) fixed; | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 129 | val _ = reject_dups "Duplicate fixed parameter(s): " fixed'; | 
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 130 | val implicit'' = | 
| 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 131 | if strict then [] | 
| 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 132 | else | 
| 55639 | 133 | let | 
| 134 | val _ = | |
| 135 | reject_dups | |
| 136 | "Parameter(s) declared simultaneously in expression and for clause: " | |
| 137 | (implicit' @ fixed'); | |
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 138 | in map (fn (x, mx) => (Binding.name x, NONE, mx)) implicit end; | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 139 | |
| 28895 | 140 | in (expr', implicit'' @ fixed) end; | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 141 | |
| 28795 | 142 | |
| 143 | (** Read instantiation **) | |
| 144 | ||
| 28872 | 145 | (* Parse positional or named instantiation *) | 
| 146 | ||
| 28859 | 147 | local | 
| 148 | ||
| 29797 | 149 | fun prep_inst prep_term ctxt parms (Positional insts) = | 
| 49817 | 150 | (insts ~~ parms) |> map | 
| 151 | (fn (NONE, p) => Free (p, dummyT) | |
| 152 | | (SOME t, _) => prep_term ctxt t) | |
| 29797 | 153 | | prep_inst prep_term ctxt parms (Named insts) = | 
| 49817 | 154 | parms |> map (fn p => | 
| 155 | (case AList.lookup (op =) insts p of | |
| 156 | SOME t => prep_term ctxt t | | |
| 157 | NONE => Free (p, dummyT))); | |
| 28872 | 158 | |
| 159 | in | |
| 160 | ||
| 161 | fun parse_inst x = prep_inst Syntax.parse_term x; | |
| 162 | fun make_inst x = prep_inst (K I) x; | |
| 163 | ||
| 164 | end; | |
| 165 | ||
| 166 | ||
| 167 | (* Instantiation morphism *) | |
| 168 | ||
| 67714 | 169 | fun inst_morphism params ((prfx, mandatory), insts') ctxt = | 
| 28795 | 170 | let | 
| 171 | (* parameters *) | |
| 67714 | 172 | val parm_types = map #2 params; | 
| 67696 | 173 | val type_parms = fold Term.add_tfreesT parm_types []; | 
| 28795 | 174 | |
| 67699 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 wenzelm parents: 
67696diff
changeset | 175 | (* type inference *) | 
| 37145 
01aa36932739
renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
 wenzelm parents: 
36674diff
changeset | 176 | val parm_types' = map (Type_Infer.paramify_vars o Logic.varifyT_global) parm_types; | 
| 67696 | 177 | val type_parms' = fold Term.add_tvarsT parm_types' []; | 
| 178 | val checked = | |
| 179 | (map (Logic.mk_type o TVar) type_parms' @ map2 Type.constraint parm_types' insts') | |
| 180 | |> Syntax.check_terms (Config.put Type_Infer.object_logic false ctxt) | |
| 181 | val (type_parms'', insts'') = chop (length type_parms') checked; | |
| 182 | ||
| 67699 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 wenzelm parents: 
67696diff
changeset | 183 | (* context *) | 
| 70308 | 184 | val ctxt' = fold Proof_Context.augment checked ctxt; | 
| 67699 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 wenzelm parents: 
67696diff
changeset | 185 | val certT = Thm.trim_context_ctyp o Thm.ctyp_of ctxt'; | 
| 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 wenzelm parents: 
67696diff
changeset | 186 | val cert = Thm.trim_context_cterm o Thm.cterm_of ctxt'; | 
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 187 | |
| 28795 | 188 | (* instantiation *) | 
| 67696 | 189 | val instT = | 
| 74282 | 190 | TFrees.build | 
| 191 | (fold2 (fn v => fn T => not (TFree v = T) ? TFrees.add (v, T)) | |
| 192 | type_parms (map Logic.dest_type type_parms'')); | |
| 67699 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 wenzelm parents: 
67696diff
changeset | 193 | val cert_inst = | 
| 74282 | 194 | Frees.build | 
| 195 | (fold2 (fn v => fn t => not (Free v = t) ? Frees.add (v, cert t)) | |
| 196 | (map #1 params ~~ map (Term_Subst.instantiateT_frees instT) parm_types) insts''); | |
| 28795 | 197 | in | 
| 74282 | 198 | (Element.instantiate_normalize_morphism (TFrees.map (K certT) instT, cert_inst) $> | 
| 54740 | 199 | Morphism.binding_morphism "Expression.inst" (Binding.prefix mandatory prfx), ctxt') | 
| 28795 | 200 | end; | 
| 28859 | 201 | |
| 28795 | 202 | |
| 203 | (*** Locale processing ***) | |
| 204 | ||
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 205 | (** Parsing **) | 
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 206 | |
| 29604 | 207 | fun parse_elem prep_typ prep_term ctxt = | 
| 208 | Element.map_ctxt | |
| 209 |    {binding = I,
 | |
| 210 | typ = prep_typ ctxt, | |
| 42360 | 211 | term = prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt), | 
| 212 | pattern = prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt), | |
| 29604 | 213 | fact = I, | 
| 214 | attrib = I}; | |
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 215 | |
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 216 | fun prepare_stmt prep_prop prep_obtains ctxt stmt = | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 217 | (case stmt of | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 218 | Element.Shows raw_shows => | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 219 | raw_shows |> (map o apsnd o map) (fn (t, ps) => | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 220 | (prep_prop (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) t, | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 221 | map (prep_prop (Proof_Context.set_mode Proof_Context.mode_pattern ctxt)) ps)) | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 222 | | Element.Obtains raw_obtains => | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 223 | let | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 224 | val ((_, thesis), thesis_ctxt) = Obtain.obtain_thesis ctxt; | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 225 | val obtains = prep_obtains thesis_ctxt thesis raw_obtains; | 
| 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 226 | in map (fn (b, t) => ((b, []), [(t, [])])) obtains end); | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 227 | |
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 228 | |
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 229 | (** Simultaneous type inference: instantiations + elements + statement **) | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 230 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 231 | local | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 232 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 233 | fun mk_type T = (Logic.mk_type T, []); | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 234 | fun mk_term t = (t, []); | 
| 39288 | 235 | fun mk_propp (p, pats) = (Type.constraint propT p, pats); | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 236 | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 237 | fun dest_type (T, []) = Logic.dest_type T; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 238 | fun dest_term (t, []) = t; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 239 | fun dest_propp (p, pats) = (p, pats); | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 240 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 241 | fun extract_inst (_, (_, ts)) = map mk_term ts; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 242 | fun restore_inst ((l, (p, _)), cs) = (l, (p, map dest_term cs)); | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 243 | |
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 244 | fun extract_eqns es = map (mk_term o snd) es; | 
| 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 245 | fun restore_eqns (es, cs) = map2 (fn (b, _) => fn c => (b, dest_term c)) es cs; | 
| 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 246 | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 247 | fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 248 | | extract_elem (Constrains csts) = map (#2 #> single #> map mk_type) csts | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 249 | | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 250 | | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs | 
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67665diff
changeset | 251 | | extract_elem (Notes _) = [] | 
| 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67665diff
changeset | 252 | | extract_elem (Lazy_Notes _) = []; | 
| 28795 | 253 | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 254 | fun restore_elem (Fixes fixes, css) = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 255 | (fixes ~~ css) |> map (fn ((x, _, mx), cs) => | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 256 | (x, cs |> map dest_type |> try hd, mx)) |> Fixes | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 257 | | restore_elem (Constrains csts, css) = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 258 | (csts ~~ css) |> map (fn ((x, _), cs) => | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 259 | (x, cs |> map dest_type |> hd)) |> Constrains | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 260 | | restore_elem (Assumes asms, css) = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 261 | (asms ~~ css) |> map (fn ((b, _), cs) => (b, map dest_propp cs)) |> Assumes | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 262 | | restore_elem (Defines defs, css) = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 263 | (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines | 
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67665diff
changeset | 264 | | restore_elem (elem as Notes _, _) = elem | 
| 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67665diff
changeset | 265 | | restore_elem (elem as Lazy_Notes _, _) = elem; | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 266 | |
| 63084 | 267 | fun prep (_, pats) (ctxt, t :: ts) = | 
| 70308 | 268 | let val ctxt' = Proof_Context.augment t ctxt | 
| 63084 | 269 | in | 
| 270 | ((t, Syntax.check_props (Proof_Context.set_mode Proof_Context.mode_pattern ctxt') pats), | |
| 271 | (ctxt', ts)) | |
| 272 | end; | |
| 273 | ||
| 274 | fun check cs ctxt = | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 275 | let | 
| 63084 | 276 | val (cs', (ctxt', _)) = fold_map prep cs | 
| 277 | (ctxt, Syntax.check_terms | |
| 278 | (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) (map fst cs)); | |
| 279 | in (cs', ctxt') end; | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 280 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 281 | in | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 282 | |
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 283 | fun check_autofix insts eqnss elems concl ctxt = | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 284 | let | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 285 | val inst_cs = map extract_inst insts; | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 286 | val eqns_cs = map extract_eqns eqnss; | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 287 | val elem_css = map extract_elem elems; | 
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 288 | val concl_cs = (map o map) mk_propp (map snd concl); | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 289 | (* Type inference *) | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 290 | val (inst_cs' :: eqns_cs' :: css', ctxt') = | 
| 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 291 | (fold_burrow o fold_burrow) check (inst_cs :: eqns_cs :: elem_css @ [concl_cs]) ctxt; | 
| 28936 
f1647bf418f5
No resolution of patterns within context statements.
 ballarin parents: 
28903diff
changeset | 292 | val (elem_css', [concl_cs']) = chop (length elem_css) css'; | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 293 | in | 
| 60476 | 294 | ((map restore_inst (insts ~~ inst_cs'), | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 295 | map restore_eqns (eqnss ~~ eqns_cs'), | 
| 30776 | 296 | map restore_elem (elems ~~ elem_css'), | 
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 297 | map fst concl ~~ concl_cs'), ctxt') | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 298 | end; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 299 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 300 | end; | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 301 | |
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 302 | |
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 303 | (** Prepare locale elements **) | 
| 28795 | 304 | |
| 60379 | 305 | fun declare_elem prep_var (Fixes fixes) ctxt = | 
| 306 | let val (vars, _) = fold_map prep_var fixes ctxt | |
| 42360 | 307 | in ctxt |> Proof_Context.add_fixes vars |> snd end | 
| 60379 | 308 | | declare_elem prep_var (Constrains csts) ctxt = | 
| 309 | ctxt |> fold_map (fn (x, T) => prep_var (Binding.name x, SOME T, NoSyn)) csts |> snd | |
| 28872 | 310 | | declare_elem _ (Assumes _) ctxt = ctxt | 
| 311 | | declare_elem _ (Defines _) ctxt = ctxt | |
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67665diff
changeset | 312 | | declare_elem _ (Notes _) ctxt = ctxt | 
| 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67665diff
changeset | 313 | | declare_elem _ (Lazy_Notes _) ctxt = ctxt; | 
| 28795 | 314 | |
| 30776 | 315 | |
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 316 | (** Finish locale elements **) | 
| 28795 | 317 | |
| 49819 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 318 | fun finish_inst ctxt (loc, (prfx, inst)) = | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 319 | let | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 320 | val thy = Proof_Context.theory_of ctxt; | 
| 67714 | 321 | val (morph, _) = inst_morphism (map #1 (Locale.params_of thy loc)) (prfx, inst) ctxt; | 
| 49819 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 322 | in (loc, morph) end; | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 323 | |
| 49818 | 324 | fun finish_fixes (parms: (string * typ) list) = map (fn (binding, _, mx) => | 
| 325 | let val x = Binding.name_of binding | |
| 326 | in (binding, AList.lookup (op =) parms x, mx) end); | |
| 49817 | 327 | |
| 328 | local | |
| 329 | ||
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 330 | fun closeup _ _ false elem = elem | 
| 49819 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 331 | | closeup (outer_ctxt, ctxt) parms true elem = | 
| 28795 | 332 | let | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 333 | (* FIXME consider closing in syntactic phase -- before type checking *) | 
| 28795 | 334 | fun close_frees t = | 
| 335 | let | |
| 336 | val rev_frees = | |
| 337 | Term.fold_aterms (fn Free (x, T) => | |
| 49819 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 338 | if Variable.is_fixed outer_ctxt x orelse AList.defined (op =) parms x then I | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 339 | else insert (op =) (x, T) | _ => I) t []; | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 340 | in fold (Logic.all o Free) rev_frees t end; | 
| 28795 | 341 | |
| 342 | fun no_binds [] = [] | |
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 343 | | no_binds _ = error "Illegal term bindings in context element"; | 
| 28795 | 344 | in | 
| 345 | (case elem of | |
| 346 | Assumes asms => Assumes (asms |> map (fn (a, propps) => | |
| 347 | (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps))) | |
| 29022 | 348 | | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) => | 
| 63395 | 349 | let val ((c, _), t') = Local_Defs.cert_def ctxt (K []) (close_frees t) | 
| 30434 | 350 | in ((Thm.def_binding_optional (Binding.name c) name, atts), (t', no_binds ps)) end)) | 
| 28795 | 351 | | e => e) | 
| 352 | end; | |
| 353 | ||
| 49819 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 354 | in | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 355 | |
| 49818 | 356 | fun finish_elem _ parms _ (Fixes fixes) = Fixes (finish_fixes parms fixes) | 
| 357 | | finish_elem _ _ _ (Constrains _) = Constrains [] | |
| 49819 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 358 | | finish_elem ctxts parms do_close (Assumes asms) = closeup ctxts parms do_close (Assumes asms) | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 359 | | finish_elem ctxts parms do_close (Defines defs) = closeup ctxts parms do_close (Defines defs) | 
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67665diff
changeset | 360 | | finish_elem _ _ _ (elem as Notes _) = elem | 
| 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67665diff
changeset | 361 | | finish_elem _ _ _ (elem as Lazy_Notes _) = elem; | 
| 28872 | 362 | |
| 49817 | 363 | end; | 
| 364 | ||
| 28795 | 365 | |
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 366 | (** Process full context statement: instantiations + elements + statement **) | 
| 28895 | 367 | |
| 368 | (* Interleave incremental parsing and type inference over entire parsed stretch. *) | |
| 369 | ||
| 28795 | 370 | local | 
| 371 | ||
| 67740 | 372 | fun abs_def ctxt = | 
| 373 | Thm.cterm_of ctxt #> Assumption.assume ctxt #> Local_Defs.abs_def_rule ctxt #> Thm.prop_of; | |
| 374 | ||
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 375 | fun prep_full_context_statement | 
| 67740 | 376 | parse_typ parse_prop prep_obtains prep_var_elem prep_inst prep_eqns prep_attr prep_var_inst prep_expr | 
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 377 |     {strict, do_close, fixed_frees} raw_import init_body raw_elems raw_stmt ctxt1 =
 | 
| 28795 | 378 | let | 
| 42360 | 379 | val thy = Proof_Context.theory_of ctxt1; | 
| 28872 | 380 | |
| 28895 | 381 | val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import); | 
| 382 | ||
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 383 | fun prep_insts_cumulative (loc, (prfx, (inst, eqns))) (i, insts, eqnss, ctxt) = | 
| 28872 | 384 | let | 
| 67714 | 385 | val params = map #1 (Locale.params_of thy loc); | 
| 386 | val inst' = prep_inst ctxt (map #1 params) inst; | |
| 387 | val parm_types' = | |
| 388 | params |> map (#2 #> Logic.varifyT_global #> | |
| 67695 | 389 | Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) #> | 
| 390 | Type_Infer.paramify_vars); | |
| 39288 | 391 | val inst'' = map2 Type.constraint parm_types' inst'; | 
| 28872 | 392 | val insts' = insts @ [(loc, (prfx, inst''))]; | 
| 67740 | 393 | val ((insts'', _, _, _), ctxt2) = check_autofix insts' [] [] [] ctxt; | 
| 394 | val inst''' = insts'' |> List.last |> snd |> snd; | |
| 395 | val (inst_morph, _) = inst_morphism params (prfx, inst''') ctxt; | |
| 67741 
d5a7f2c54655
Fall back to reading rewrite morphism first if activation fails without it.
 ballarin parents: 
67740diff
changeset | 396 | val ctxt' = Locale.activate_declarations (loc, inst_morph) ctxt2 | 
| 
d5a7f2c54655
Fall back to reading rewrite morphism first if activation fails without it.
 ballarin parents: 
67740diff
changeset | 397 | handle ERROR msg => if null eqns then error msg else | 
| 70622 
2fb2e7661e16
Integrate locale activation fallback diagnostics with 'trace_locales'.
 ballarin parents: 
70314diff
changeset | 398 | (Locale.tracing ctxt1 | 
| 
2fb2e7661e16
Integrate locale activation fallback diagnostics with 'trace_locales'.
 ballarin parents: 
70314diff
changeset | 399 | (msg ^ "\nFalling back to reading rewrites clause before activation."); | 
| 
2fb2e7661e16
Integrate locale activation fallback diagnostics with 'trace_locales'.
 ballarin parents: 
70314diff
changeset | 400 | ctxt2); | 
| 67740 | 401 | |
| 402 | val attrss = map (apsnd (map (prep_attr ctxt)) o fst) eqns; | |
| 403 | val eqns' = (prep_eqns ctxt' o map snd) eqns; | |
| 404 | val eqnss' = [attrss ~~ eqns']; | |
| 405 | val ((_, [eqns''], _, _), _) = check_autofix insts'' eqnss' [] [] ctxt'; | |
| 406 | val rewrite_morph = eqns' | |
| 407 | |> map (abs_def ctxt') | |
| 408 | |> Variable.export_terms ctxt' ctxt | |
| 78453 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 wenzelm parents: 
78095diff
changeset | 409 | |> Element.eq_term_morphism ctxt | 
| 78060 | 410 | |> Morphism.default; | 
| 67742 
6306bd582957
Drop illegitimate optimisation from d5a7f2c54655.
 ballarin parents: 
67741diff
changeset | 411 | val ctxt'' = Locale.activate_declarations (loc, inst_morph $> rewrite_morph) ctxt; | 
| 67740 | 412 | val eqnss' = eqnss @ [attrss ~~ Variable.export_terms ctxt' ctxt eqns']; | 
| 413 | in (i + 1, insts', eqnss', ctxt'') end; | |
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 414 | |
| 51729 | 415 | fun prep_elem raw_elem ctxt = | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 416 | let | 
| 47315 | 417 | val ctxt' = ctxt | 
| 418 | |> Context_Position.set_visible false | |
| 60379 | 419 | |> declare_elem prep_var_elem raw_elem | 
| 47315 | 420 | |> Context_Position.restore_visible ctxt; | 
| 421 | val elems' = parse_elem parse_typ parse_prop ctxt' raw_elem; | |
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 422 | in (elems', ctxt') end; | 
| 28795 | 423 | |
| 60379 | 424 | val fors = fold_map prep_var_inst fixed ctxt1 |> fst; | 
| 42360 | 425 | val ctxt2 = ctxt1 |> Proof_Context.add_fixes fors |> snd; | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 426 | val (_, insts', eqnss', ctxt3) = fold prep_insts_cumulative raw_insts (0, [], [], ctxt2); | 
| 30786 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 427 | |
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 428 | fun prep_stmt elems ctxt = | 
| 67740 | 429 | check_autofix insts' [] elems (prepare_stmt parse_prop prep_obtains ctxt raw_stmt) ctxt; | 
| 60476 | 430 | |
| 30786 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 431 | val _ = | 
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 432 | if fixed_frees then () | 
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 433 | else | 
| 42482 | 434 | (case fold (fold (Variable.add_frees ctxt3) o snd o snd) insts' [] of | 
| 30786 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 435 | [] => () | 
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 436 |         | frees => error ("Illegal free variables in expression: " ^
 | 
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 437 | commas_quote (map (Syntax.string_of_term ctxt3 o Free) (rev frees)))); | 
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 438 | |
| 67740 | 439 | val ((insts, _, elems', concl), ctxt4) = ctxt3 | 
| 60476 | 440 | |> init_body | 
| 441 | |> fold_map prep_elem raw_elems | |
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 442 | |-> prep_stmt; | 
| 28795 | 443 | |
| 60407 | 444 | |
| 445 | (* parameters from expression and elements *) | |
| 446 | ||
| 54878 | 447 | val xs = maps (fn Fixes fixes => map (Variable.check_name o #1) fixes | _ => []) | 
| 448 | (Fixes fors :: elems'); | |
| 60476 | 449 | val (parms, ctxt5) = fold_map Proof_Context.inferred_param xs ctxt4; | 
| 28795 | 450 | |
| 49818 | 451 | val fors' = finish_fixes parms fors; | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 452 | val fixed = map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fors'; | 
| 60476 | 453 | val deps = map (finish_inst ctxt5) insts; | 
| 454 | val elems'' = map (finish_elem (ctxt1, ctxt5) parms do_close) elems'; | |
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 455 | |
| 67740 | 456 | in ((fixed, deps, eqnss', elems'', concl), (parms, ctxt5)) end; | 
| 28795 | 457 | |
| 458 | in | |
| 459 | ||
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 460 | fun cert_full_context_statement x = | 
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 461 | prep_full_context_statement (K I) (K I) Obtain.cert_obtains | 
| 67740 | 462 | Proof_Context.cert_var make_inst Syntax.check_props (K I) Proof_Context.cert_var (K I) x; | 
| 30776 | 463 | |
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 464 | fun cert_read_full_context_statement x = | 
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 465 | prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains | 
| 67740 | 466 | Proof_Context.read_var make_inst Syntax.check_props (K I) Proof_Context.cert_var (K I) x; | 
| 30776 | 467 | |
| 28895 | 468 | fun read_full_context_statement x = | 
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 469 | prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains | 
| 67740 | 470 | Proof_Context.read_var parse_inst Syntax.read_props Attrib.check_src Proof_Context.read_var check_expr x; | 
| 28795 | 471 | |
| 472 | end; | |
| 473 | ||
| 474 | ||
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 475 | (* Context statement: elements + statement *) | 
| 28795 | 476 | |
| 477 | local | |
| 478 | ||
| 63084 | 479 | fun prep_statement prep activate raw_elems raw_stmt ctxt = | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 480 | let | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 481 | val ((_, _, _, elems, concl), _) = | 
| 49817 | 482 |       prep {strict = true, do_close = false, fixed_frees = true}
 | 
| 63084 | 483 | ([], []) I raw_elems raw_stmt ctxt; | 
| 63086 
5c8e6a751adc
avoid spurious fact index, notably in "context begin" (via Bundle.context);
 wenzelm parents: 
63085diff
changeset | 484 | val ctxt' = ctxt | 
| 49817 | 485 | |> Proof_Context.set_stmt true | 
| 63086 
5c8e6a751adc
avoid spurious fact index, notably in "context begin" (via Bundle.context);
 wenzelm parents: 
63085diff
changeset | 486 | |> fold_map activate elems |> #2 | 
| 
5c8e6a751adc
avoid spurious fact index, notably in "context begin" (via Bundle.context);
 wenzelm parents: 
63085diff
changeset | 487 | |> Proof_Context.restore_stmt ctxt; | 
| 63084 | 488 | in (concl, ctxt') end; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 489 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 490 | in | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 491 | |
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 492 | fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 493 | fun read_statement x = prep_statement read_full_context_statement Element.activate x; | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 494 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 495 | end; | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 496 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 497 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 498 | (* Locale declaration: import + elements *) | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 499 | |
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 500 | fun fix_params params = | 
| 42360 | 501 | Proof_Context.add_fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params) #> snd; | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 502 | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 503 | local | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 504 | |
| 63085 | 505 | fun prep_declaration prep activate raw_import init_body raw_elems ctxt = | 
| 28795 | 506 | let | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 507 | val ((fixed, deps, eqnss, elems, _), (parms, ctxt0)) = | 
| 30786 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 508 |       prep {strict = false, do_close = true, fixed_frees = false}
 | 
| 63085 | 509 | raw_import init_body raw_elems (Element.Shows []) ctxt; | 
| 67740 | 510 | val _ = null (flat eqnss) orelse error "Illegal rewrites clause(s) in declaration of locale"; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 511 | (* Declare parameters and imported facts *) | 
| 63085 | 512 | val ctxt' = ctxt | 
| 513 | |> fix_params fixed | |
| 514 | |> fold (Context.proof_map o Locale.activate_facts NONE) deps; | |
| 515 | val (elems', ctxt'') = ctxt' | |
| 516 | |> Proof_Context.set_stmt true | |
| 63086 
5c8e6a751adc
avoid spurious fact index, notably in "context begin" (via Bundle.context);
 wenzelm parents: 
63085diff
changeset | 517 | |> fold_map activate elems | 
| 
5c8e6a751adc
avoid spurious fact index, notably in "context begin" (via Bundle.context);
 wenzelm parents: 
63085diff
changeset | 518 | ||> Proof_Context.restore_stmt ctxt'; | 
| 63085 | 519 | in ((fixed, deps, elems', ctxt''), (parms, ctxt0)) end; | 
| 28795 | 520 | |
| 521 | in | |
| 522 | ||
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 523 | fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x; | 
| 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 524 | fun cert_read_declaration x = prep_declaration cert_read_full_context_statement Element.activate x; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 525 | fun read_declaration x = prep_declaration read_full_context_statement Element.activate x; | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 526 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 527 | end; | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 528 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 529 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 530 | (* Locale expression to set up a goal *) | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 531 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 532 | local | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 533 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 534 | fun props_of thy (name, morph) = | 
| 68861 | 535 | let val (asm, defs) = Locale.specification_of thy name | 
| 536 | in map (Morphism.term morph) (the_list asm @ defs) end; | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 537 | |
| 63084 | 538 | fun prep_goal_expression prep expression ctxt = | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 539 | let | 
| 63084 | 540 | val thy = Proof_Context.theory_of ctxt; | 
| 28879 | 541 | |
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 542 | val ((fixed, deps, eqnss, _, _), _) = | 
| 60477 
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
 wenzelm parents: 
60476diff
changeset | 543 |       prep {strict = true, do_close = true, fixed_frees = true} expression I []
 | 
| 63084 | 544 | (Element.Shows []) ctxt; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 545 | (* proof obligations *) | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 546 | val propss = map (props_of thy) deps; | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 547 | val eq_propss = (map o map) snd eqnss; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 548 | |
| 63084 | 549 | val goal_ctxt = ctxt | 
| 550 | |> fix_params fixed | |
| 70308 | 551 | |> (fold o fold) Proof_Context.augment (propss @ eq_propss); | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 552 | |
| 70314 
6d6839a948cf
proper Proof_Context.export_morphism corresponding to Proof_Context.augment (see 7f568724d67e);
 wenzelm parents: 
70308diff
changeset | 553 | val export = Proof_Context.export_morphism goal_ctxt ctxt; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 554 | val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export; | 
| 31977 | 555 | val exp_term = Term_Subst.zero_var_indexes o Morphism.term export; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 556 | val exp_typ = Logic.type_map exp_term; | 
| 45289 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 wenzelm parents: 
43543diff
changeset | 557 | val export' = | 
| 54740 | 558 | Morphism.morphism "Expression.prep_goal" | 
| 78062 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78060diff
changeset | 559 |         {binding = [], typ = [K exp_typ], term = [K exp_term], fact = [K exp_fact]};
 | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 560 | in ((propss, eq_propss, deps, eqnss, export'), goal_ctxt) end; | 
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 561 | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 562 | in | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 563 | |
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 564 | fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 565 | fun read_goal_expression x = prep_goal_expression read_full_context_statement x; | 
| 28879 | 566 | |
| 28795 | 567 | end; | 
| 568 | ||
| 569 | ||
| 570 | (*** Locale declarations ***) | |
| 571 | ||
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 572 | (* extract specification text *) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 573 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 574 | val norm_term = Envir.beta_norm oo Term.subst_atomic; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 575 | |
| 74275 
aed955bb4cb1
omit obsolete field "xs": originally from fd0f8fa2b6bd, but later unused;
 wenzelm parents: 
74266diff
changeset | 576 | fun bind_def ctxt eq (env, eqs) = | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 577 | let | 
| 63395 | 578 | val _ = Local_Defs.cert_def ctxt (K []) eq; | 
| 35624 | 579 | val ((y, T), b) = Local_Defs.abs_def eq; | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 580 | val b' = norm_term env b; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 581 | fun err msg = error (msg ^ ": " ^ quote y); | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 582 | in | 
| 49749 | 583 | (case filter (fn (Free (y', _), _) => y = y' | _ => false) env of | 
| 74275 
aed955bb4cb1
omit obsolete field "xs": originally from fd0f8fa2b6bd, but later unused;
 wenzelm parents: 
74266diff
changeset | 584 | [] => ((Free (y, T), b') :: env, eq :: eqs) | 
| 49749 | 585 | | dups => | 
| 74275 
aed955bb4cb1
omit obsolete field "xs": originally from fd0f8fa2b6bd, but later unused;
 wenzelm parents: 
74266diff
changeset | 586 | if forall (fn (_, b'') => b' aconv b'') dups then (env, eqs) | 
| 49749 | 587 | else err "Attempt to redefine variable") | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 588 | end; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 589 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 590 | (* text has the following structure: | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 591 | (((exts, exts'), (ints, ints')), (xs, env, defs)) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 592 | where | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 593 | exts: external assumptions (terms in assumes elements) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 594 | exts': dito, normalised wrt. env | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 595 | ints: internal assumptions (terms in assumptions from insts) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 596 | ints': dito, normalised wrt. env | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 597 | xs: the free variables in exts' and ints' and rhss of definitions, | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 598 | this includes parameters except defined parameters | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 599 | env: list of term pairs encoding substitutions, where the first term | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 600 | is a free variable; substitutions represent defines elements and | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 601 | the rhs is normalised wrt. the previous env | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 602 | defs: the equations from the defines elements | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 603 | *) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 604 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 605 | fun eval_text _ _ (Fixes _) text = text | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 606 | | eval_text _ _ (Constrains _) text = text | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 607 | | eval_text _ is_ext (Assumes asms) | 
| 74275 
aed955bb4cb1
omit obsolete field "xs": originally from fd0f8fa2b6bd, but later unused;
 wenzelm parents: 
74266diff
changeset | 608 | (((exts, exts'), (ints, ints')), (env, defs)) = | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 609 | let | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 610 | val ts = maps (map #1 o #2) asms; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 611 | val ts' = map (norm_term env) ts; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 612 | val spec' = | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 613 | if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints')) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 614 | else ((exts, exts'), (ints @ ts, ints' @ ts')); | 
| 74275 
aed955bb4cb1
omit obsolete field "xs": originally from fd0f8fa2b6bd, but later unused;
 wenzelm parents: 
74266diff
changeset | 615 | in (spec', (env, defs)) end | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 616 | | eval_text ctxt _ (Defines defs) (spec, binds) = | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 617 | (spec, fold (bind_def ctxt o #1 o #2) defs binds) | 
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67665diff
changeset | 618 | | eval_text _ _ (Notes _) text = text | 
| 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67665diff
changeset | 619 | | eval_text _ _ (Lazy_Notes _) text = text; | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 620 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 621 | fun eval_inst ctxt (loc, morph) text = | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 622 | let | 
| 42360 | 623 | val thy = Proof_Context.theory_of ctxt; | 
| 29360 | 624 | val (asm, defs) = Locale.specification_of thy loc; | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 625 | val asm' = Option.map (Morphism.term morph) asm; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 626 | val defs' = map (Morphism.term morph) defs; | 
| 55639 | 627 | val text' = | 
| 628 | text |> | |
| 629 | (if is_some asm then | |
| 63352 | 630 | eval_text ctxt false (Assumes [(Binding.empty_atts, [(the asm', [])])]) | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 631 | else I) |> | 
| 55639 | 632 | (if not (null defs) then | 
| 63352 | 633 | eval_text ctxt false (Defines (map (fn def => (Binding.empty_atts, (def, []))) defs')) | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 634 | else I) | 
| 29360 | 635 | (* FIXME clone from locale.ML *) | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 636 | in text' end; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 637 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 638 | fun eval_elem ctxt elem text = | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 639 | eval_text ctxt true elem text; | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 640 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 641 | fun eval ctxt deps elems = | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 642 | let | 
| 74275 
aed955bb4cb1
omit obsolete field "xs": originally from fd0f8fa2b6bd, but later unused;
 wenzelm parents: 
74266diff
changeset | 643 | val text' = fold (eval_inst ctxt) deps ((([], []), ([], [])), ([], [])); | 
| 
aed955bb4cb1
omit obsolete field "xs": originally from fd0f8fa2b6bd, but later unused;
 wenzelm parents: 
74266diff
changeset | 644 | val ((spec, (_, defs))) = fold (eval_elem ctxt) elems text'; | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 645 | in (spec, defs) end; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 646 | |
| 28903 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 647 | (* axiomsN: name of theorem set with destruct rules for locale predicates, | 
| 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 648 | also name suffix of delta predicates and assumptions. *) | 
| 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 649 | |
| 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 650 | val axiomsN = "axioms"; | 
| 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 651 | |
| 28795 | 652 | local | 
| 653 | ||
| 654 | (* introN: name of theorems for introduction rules of locale and | |
| 28903 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 655 | delta predicates *) | 
| 28795 | 656 | |
| 657 | val introN = "intro"; | |
| 658 | ||
| 59970 | 659 | fun atomize_spec ctxt ts = | 
| 28795 | 660 | let | 
| 661 | val t = Logic.mk_conjunction_balanced ts; | |
| 59970 | 662 | val body = Object_Logic.atomize_term ctxt t; | 
| 28795 | 663 | val bodyT = Term.fastype_of body; | 
| 664 | in | |
| 55639 | 665 | if bodyT = propT | 
| 59970 | 666 | then (t, propT, Thm.reflexive (Thm.cterm_of ctxt t)) | 
| 667 | else (body, bodyT, Object_Logic.atomize ctxt (Thm.cterm_of ctxt t)) | |
| 28795 | 668 | end; | 
| 669 | ||
| 670 | (* achieve plain syntax for locale predicates (without "PROP") *) | |
| 671 | ||
| 49820 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 672 | fun aprop_tr' n c = | 
| 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 673 | let | 
| 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 674 | val c' = Lexicon.mark_const c; | 
| 52143 | 675 | fun tr' (_: Proof.context) T args = | 
| 49820 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 676 | if T <> dummyT andalso length args = n | 
| 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 677 | then Syntax.const "_aprop" $ Term.list_comb (Syntax.const c', args) | 
| 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 678 | else raise Match; | 
| 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 679 | in (c', tr') end; | 
| 28795 | 680 | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 681 | (* define one predicate including its intro rule and axioms | 
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 682 | - binding: predicate name | 
| 28795 | 683 | - parms: locale parameters | 
| 684 | - defs: thms representing substitutions from defines elements | |
| 685 | - ts: terms representing locale assumptions (not normalised wrt. defs) | |
| 686 | - norm_ts: terms representing locale assumptions (normalised wrt. defs) | |
| 687 | - thy: the theory | |
| 688 | *) | |
| 689 | ||
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 690 | fun def_pred binding parms defs ts norm_ts thy = | 
| 28795 | 691 | let | 
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 692 | val name = Sign.full_name thy binding; | 
| 28795 | 693 | |
| 59970 | 694 | val thy_ctxt = Proof_Context.init_global thy; | 
| 695 | ||
| 696 | val (body, bodyT, body_eq) = atomize_spec thy_ctxt norm_ts; | |
| 74276 | 697 | val env = Names.build (Names.add_free_names body); | 
| 698 | val xs = filter (Names.defined env o #1) parms; | |
| 28795 | 699 | val Ts = map #2 xs; | 
| 74276 | 700 | val type_tfrees = TFrees.build (fold TFrees.add_tfreesT Ts); | 
| 701 | val extra_tfrees = | |
| 74278 | 702 | TFrees.build (TFrees.add_tfrees_unless (TFrees.defined type_tfrees) body) | 
| 74279 
42db84eaee2d
clarified order of extra type variables, following names more often than occurrences;
 wenzelm parents: 
74278diff
changeset | 703 | |> TFrees.keys |> map TFree; | 
| 74276 | 704 | val predT = map Term.itselfT extra_tfrees ---> Ts ---> bodyT; | 
| 28795 | 705 | |
| 74276 | 706 | val args = map Logic.mk_type extra_tfrees @ map Free xs; | 
| 28795 | 707 | val head = Term.list_comb (Const (name, predT), args); | 
| 59970 | 708 | val statement = Object_Logic.ensure_propT thy_ctxt head; | 
| 28795 | 709 | |
| 79336 
032a31db4c6f
clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
 wenzelm parents: 
78453diff
changeset | 710 | val (pred_def, defs_thy) = | 
| 28795 | 711 | thy | 
| 52143 | 712 | |> bodyT = propT ? Sign.typed_print_translation [aprop_tr' (length args) name] | 
| 61082 
42c2698d2273
expose locale definition to normal user-namespace (for completion, query etc.) -- in contrast to 149f80f27c84, ba9f52f56356, f7d9c5e5d2f9;
 wenzelm parents: 
60949diff
changeset | 713 | |> Sign.declare_const_global ((binding, predT), NoSyn) |> snd | 
| 79336 
032a31db4c6f
clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
 wenzelm parents: 
78453diff
changeset | 714 | |> Global_Theory.add_def (Thm.def_binding binding, Logic.mk_equals (head, body)); | 
| 42360 | 715 | val defs_ctxt = Proof_Context.init_global defs_thy |> Variable.declare_term head; | 
| 28795 | 716 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54740diff
changeset | 717 | val intro = Goal.prove_global defs_thy [] norm_ts statement | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54740diff
changeset | 718 |       (fn {context = ctxt, ...} =>
 | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54740diff
changeset | 719 | rewrite_goals_tac ctxt [pred_def] THEN | 
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
57926diff
changeset | 720 | compose_tac defs_ctxt (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN | 
| 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
57926diff
changeset | 721 | compose_tac defs_ctxt | 
| 59616 | 722 | (false, | 
| 71018 
d32ed8927a42
more robust expose_proofs corresponding to register_proofs/consolidate_theory;
 wenzelm parents: 
71017diff
changeset | 723 | Conjunction.intr_balanced (map (Thm.assume o Thm.cterm_of defs_ctxt) norm_ts), 0) 1); | 
| 28795 | 724 | |
| 725 | val conjuncts = | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54740diff
changeset | 726 | (Drule.equal_elim_rule2 OF | 
| 59623 | 727 | [body_eq, rewrite_rule defs_ctxt [pred_def] (Thm.assume (Thm.cterm_of defs_ctxt statement))]) | 
| 28795 | 728 | |> Conjunction.elim_balanced (length ts); | 
| 54566 | 729 | |
| 730 | val (_, axioms_ctxt) = defs_ctxt | |
| 60949 | 731 | |> Assumption.add_assumes (maps Thm.chyps_of (defs @ conjuncts)); | 
| 28795 | 732 | val axioms = ts ~~ conjuncts |> map (fn (t, ax) => | 
| 54566 | 733 | Element.prove_witness axioms_ctxt t | 
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
57926diff
changeset | 734 | (rewrite_goals_tac axioms_ctxt defs THEN compose_tac axioms_ctxt (false, ax, 0) 1)); | 
| 28795 | 735 | in ((statement, intro, axioms), defs_thy) end; | 
| 736 | ||
| 737 | in | |
| 738 | ||
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 739 | (* main predicate definition function *) | 
| 28795 | 740 | |
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 741 | fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy = | 
| 28795 | 742 | let | 
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54879diff
changeset | 743 | val ctxt = Proof_Context.init_global thy; | 
| 59623 | 744 | val defs' = map (Thm.cterm_of ctxt #> Assumption.assume ctxt #> Drule.abs_def) defs; | 
| 29031 
e74341997a48
Pass on defines in inheritance; reject illicit defines created by instantiation.
 ballarin parents: 
29030diff
changeset | 745 | |
| 78042 | 746 | val (a_pred, a_intro, a_axioms, thy2) = | 
| 28795 | 747 | if null exts then (NONE, NONE, [], thy) | 
| 748 | else | |
| 749 | let | |
| 55639 | 750 | val abinding = | 
| 751 |             if null ints then binding else Binding.suffix_name ("_" ^ axiomsN) binding;
 | |
| 78042 | 752 | val ((statement, intro, axioms), thy1) = | 
| 28795 | 753 | thy | 
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 754 | |> def_pred abinding parms defs' exts exts'; | 
| 78042 | 755 | val ((_, [intro']), thy2) = | 
| 756 | thy1 | |
| 35204 
214ec053128e
locale: more precise treatment of naming vs. binding;
 wenzelm parents: 
35143diff
changeset | 757 | |> Sign.qualified_path true abinding | 
| 67713 | 758 | |> Global_Theory.note_thms "" | 
| 759 | ((Binding.name introN, []), [([intro], [Locale.unfold_add])]) | |
| 78042 | 760 | ||> Sign.restore_naming thy1; | 
| 761 | in (SOME statement, SOME intro', axioms, thy2) end; | |
| 762 | val (b_pred, b_intro, b_axioms, thy4) = | |
| 763 | if null ints then (NONE, NONE, [], thy2) | |
| 28795 | 764 | else | 
| 765 | let | |
| 78042 | 766 | val ((statement, intro, axioms), thy3) = | 
| 767 | thy2 | |
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 768 | |> def_pred binding parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred); | 
| 78041 | 769 | val conclude_witness = | 
| 78042 | 770 | Drule.export_without_context o Element.conclude_witness (Proof_Context.init_global thy3); | 
| 771 | val ([(_, [intro']), _], thy4) = | |
| 772 | thy3 | |
| 35204 
214ec053128e
locale: more precise treatment of naming vs. binding;
 wenzelm parents: 
35143diff
changeset | 773 | |> Sign.qualified_path true binding | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39288diff
changeset | 774 | |> Global_Theory.note_thmss "" | 
| 61082 
42c2698d2273
expose locale definition to normal user-namespace (for completion, query etc.) -- in contrast to 149f80f27c84, ba9f52f56356, f7d9c5e5d2f9;
 wenzelm parents: 
60949diff
changeset | 775 | [((Binding.name introN, []), [([intro], [Locale.intro_add])]), | 
| 
42c2698d2273
expose locale definition to normal user-namespace (for completion, query etc.) -- in contrast to 149f80f27c84, ba9f52f56356, f7d9c5e5d2f9;
 wenzelm parents: 
60949diff
changeset | 776 | ((Binding.name axiomsN, []), | 
| 78041 | 777 | [(map conclude_witness axioms, [])])] | 
| 78042 | 778 | ||> Sign.restore_naming thy3; | 
| 779 | in (SOME statement, SOME intro', axioms, thy4) end; | |
| 780 | in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy4) end; | |
| 28795 | 781 | |
| 782 | end; | |
| 783 | ||
| 784 | ||
| 785 | local | |
| 786 | ||
| 787 | fun assumes_to_notes (Assumes asms) axms = | |
| 788 | fold_map (fn (a, spec) => fn axs => | |
| 789 | let val (ps, qs) = chop (length spec) axs | |
| 790 | in ((a, [(ps, [])]), qs) end) asms axms | |
| 33644 
5266a72e0889
eliminated slightly odd (unused) "axiom" and "assumption" -- collapsed to unspecific "";
 wenzelm parents: 
33643diff
changeset | 791 | |> apfst (curry Notes "") | 
| 28795 | 792 | | assumes_to_notes e axms = (e, axms); | 
| 793 | ||
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54879diff
changeset | 794 | fun defines_to_notes ctxt (Defines defs) = | 
| 42440 
5e7a7343ab11
discontinuend obsolete Thm.definitionK, which was hardly ever well-defined;
 wenzelm parents: 
42381diff
changeset | 795 |       Notes ("", map (fn (a, (def, _)) =>
 | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59616diff
changeset | 796 | (a, [([Assumption.assume ctxt (Thm.cterm_of ctxt def)], | 
| 78095 | 797 |           [Attrib.internal @{here} (K Locale.witness_add)])])) defs)
 | 
| 29031 
e74341997a48
Pass on defines in inheritance; reject illicit defines created by instantiation.
 ballarin parents: 
29030diff
changeset | 798 | | defines_to_notes _ e = e; | 
| 28795 | 799 | |
| 67665 | 800 | val is_hyp = fn Assumes _ => true | Defines _ => true | _ => false; | 
| 59296 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 haftmann parents: 
58956diff
changeset | 801 | |
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 802 | fun gen_add_locale prep_include prep_decl | 
| 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 803 | binding raw_predicate_binding raw_includes raw_import raw_body thy = | 
| 28795 | 804 | let | 
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 805 | val name = Sign.full_name thy binding; | 
| 29391 
1f6e8c00dc3e
tuned signature; changed locale predicate name convention
 haftmann parents: 
29362diff
changeset | 806 | val _ = Locale.defined thy name andalso | 
| 28795 | 807 |       error ("Duplicate definition of locale " ^ quote name);
 | 
| 808 | ||
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 809 | val ctxt = Proof_Context.init_global thy; | 
| 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 810 | val includes = map (prep_include ctxt) raw_includes; | 
| 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 811 | |
| 47311 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 wenzelm parents: 
47287diff
changeset | 812 | val ((fixed, deps, body_elems, _), (parms, ctxt')) = | 
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 813 | ctxt | 
| 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 814 | |> Bundle.includes includes | 
| 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 815 | |> prep_decl raw_import I raw_body; | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 816 | val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 817 | |
| 74277 | 818 | val type_tfrees = TFrees.build (fold (TFrees.add_tfreesT o #2) parms); | 
| 74276 | 819 | val extra_tfrees = | 
| 74278 | 820 | TFrees.build (fold (TFrees.add_tfrees_unless (TFrees.defined type_tfrees)) exts') | 
| 74279 
42db84eaee2d
clarified order of extra type variables, following names more often than occurrences;
 wenzelm parents: 
74278diff
changeset | 821 | |> TFrees.keys; | 
| 37313 | 822 | val _ = | 
| 74279 
42db84eaee2d
clarified order of extra type variables, following names more often than occurrences;
 wenzelm parents: 
74278diff
changeset | 823 | if null extra_tfrees then () | 
| 37313 | 824 |       else warning ("Additional type variable(s) in locale specification " ^
 | 
| 42381 
309ec68442c6
added Binding.print convenience, which includes quote already;
 wenzelm parents: 
42375diff
changeset | 825 | Binding.print binding ^ ": " ^ | 
| 74279 
42db84eaee2d
clarified order of extra type variables, following names more often than occurrences;
 wenzelm parents: 
74278diff
changeset | 826 | commas (map (Syntax.string_of_typ ctxt' o TFree) extra_tfrees)); | 
| 37313 | 827 | |
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 828 | val predicate_binding = | 
| 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 829 | if Binding.is_empty raw_predicate_binding then binding | 
| 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 830 | else raw_predicate_binding; | 
| 28872 | 831 | val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') = | 
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 832 | define_preds predicate_binding parms text thy; | 
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54879diff
changeset | 833 | val pred_ctxt = Proof_Context.init_global thy'; | 
| 28795 | 834 | |
| 29035 | 835 | val a_satisfy = Element.satisfy_morphism a_axioms; | 
| 836 | val b_satisfy = Element.satisfy_morphism b_axioms; | |
| 28903 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 837 | |
| 28895 | 838 | val params = fixed @ | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 839 | maps (fn Fixes fixes => | 
| 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 840 | map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fixes | _ => []) body_elems; | 
| 28903 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 841 | val asm = if is_some b_statement then b_statement else a_statement; | 
| 29028 
b5dad96c755a
When adding locales, delay notes until local theory is built.
 ballarin parents: 
29022diff
changeset | 842 | |
| 59296 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 haftmann parents: 
58956diff
changeset | 843 | val hyp_spec = filter is_hyp body_elems; | 
| 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 haftmann parents: 
58956diff
changeset | 844 | |
| 29028 
b5dad96c755a
When adding locales, delay notes until local theory is built.
 ballarin parents: 
29022diff
changeset | 845 | val notes = | 
| 33278 | 846 | if is_some asm then | 
| 61082 
42c2698d2273
expose locale definition to normal user-namespace (for completion, query etc.) -- in contrast to 149f80f27c84, ba9f52f56356, f7d9c5e5d2f9;
 wenzelm parents: 
60949diff
changeset | 847 |         [("", [((Binding.suffix_name ("_" ^ axiomsN) binding, []),
 | 
| 59623 | 848 | [([Assumption.assume pred_ctxt (Thm.cterm_of pred_ctxt (the asm))], | 
| 78095 | 849 |             [Attrib.internal @{here} (K Locale.witness_add)])])])]
 | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 850 | else []; | 
| 28795 | 851 | |
| 55639 | 852 | val notes' = | 
| 853 | body_elems | |
| 78064 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 wenzelm parents: 
78062diff
changeset | 854 | |> map (Element.transfer_ctxt thy') | 
| 55639 | 855 | |> map (defines_to_notes pred_ctxt) | 
| 856 | |> map (Element.transform_ctxt a_satisfy) | |
| 857 | |> (fn elems => | |
| 858 | fold_map assumes_to_notes elems (map (Element.conclude_witness pred_ctxt) a_axioms)) | |
| 859 | |> fst | |
| 860 | |> map (Element.transform_ctxt b_satisfy) | |
| 861 | |> map_filter (fn Notes notes => SOME notes | _ => NONE); | |
| 29035 | 862 | |
| 863 | val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps; | |
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54879diff
changeset | 864 | val axioms = map (Element.conclude_witness pred_ctxt) b_axioms; | 
| 28872 | 865 | |
| 29358 | 866 | val loc_ctxt = thy' | 
| 74279 
42db84eaee2d
clarified order of extra type variables, following names more often than occurrences;
 wenzelm parents: 
74278diff
changeset | 867 | |> Locale.register_locale binding (extra_tfrees, params) | 
| 59296 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 haftmann parents: 
58956diff
changeset | 868 | (asm, rev defs) (a_intro, b_intro) axioms hyp_spec [] (rev notes) (rev deps') | 
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 869 | |> Named_Target.init includes name | 
| 33671 | 870 | |> fold (fn (kind, facts) => Local_Theory.notes_kind kind facts #> snd) notes'; | 
| 29028 
b5dad96c755a
When adding locales, delay notes until local theory is built.
 ballarin parents: 
29022diff
changeset | 871 | |
| 29358 | 872 | in (name, loc_ctxt) end; | 
| 28795 | 873 | |
| 874 | in | |
| 875 | ||
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 876 | val add_locale = gen_add_locale (K I) cert_declaration; | 
| 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
71166diff
changeset | 877 | val add_locale_cmd = gen_add_locale Bundle.check read_declaration; | 
| 28795 | 878 | |
| 879 | end; | |
| 880 | ||
| 28993 
829e684b02ef
Interpretation in theories including interaction with subclass relation.
 ballarin parents: 
28951diff
changeset | 881 | end; |