| author | wenzelm | 
| Thu, 18 Oct 2012 14:15:46 +0200 | |
| changeset 49912 | c6307ee2665d | 
| parent 49820 | f7a1e1745b7b | 
| child 51565 | 5e9fdbdf88ce | 
| 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 | 
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 11 |   type ('name, 'term) expr = ('name * ((string * bool) * 'term map)) list
 | 
| 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 12 | type expression_i = (string, term) expr * (binding * typ option * mixfix) list | 
| 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 13 | 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 | 14 | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 15 | (* Processing of context statements *) | 
| 29441 | 16 | val cert_statement: Element.context_i list -> (term * term list) list list -> | 
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 17 | Proof.context -> (term * term list) list list * Proof.context | 
| 28879 | 18 | val read_statement: Element.context list -> (string * string list) list list -> | 
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 19 | Proof.context -> (term * term list) list list * Proof.context | 
| 28879 | 20 | |
| 28795 | 21 | (* Declaring locales *) | 
| 29702 | 22 | val cert_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context_i list -> | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 23 | Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list | 
| 47311 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 wenzelm parents: 
47287diff
changeset | 24 | * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) | 
| 29702 | 25 | val cert_read_declaration: expression_i -> (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 | 26 | Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list | 
| 47311 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 wenzelm parents: 
47287diff
changeset | 27 | * 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 | 28 | (*FIXME*) | 
| 29702 | 29 | 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 | 30 | Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list | 
| 47311 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 wenzelm parents: 
47287diff
changeset | 31 | * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) | 
| 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 | 32 | val add_locale: (local_theory -> local_theory) -> binding -> binding -> | 
| 
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 | 33 | expression_i -> Element.context_i list -> theory -> string * local_theory | 
| 
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 | 34 | val add_locale_cmd: (local_theory -> local_theory) -> binding -> binding -> | 
| 
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 | 35 | expression -> Element.context list -> theory -> string * local_theory | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 36 | |
| 28895 | 37 | (* Interpretation *) | 
| 29441 | 38 | val cert_goal_expression: expression_i -> Proof.context -> | 
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 39 | (term list list * (string * morphism) list * morphism) * Proof.context | 
| 29496 | 40 | val read_goal_expression: expression -> Proof.context -> | 
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 41 | (term list list * (string * morphism) list * morphism) * Proof.context | 
| 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 | 42 | val sublocale: (local_theory -> local_theory) -> string -> expression_i -> | 
| 
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 | 43 | (Attrib.binding * term) list -> theory -> Proof.state | 
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 44 | val sublocale_cmd: (local_theory -> local_theory) -> xstring * Position.T -> expression -> | 
| 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 | 45 | (Attrib.binding * string) list -> theory -> Proof.state | 
| 47287 | 46 | val interpretation: expression_i -> (Attrib.binding * term) list -> theory -> Proof.state | 
| 47 | val interpretation_cmd: expression -> (Attrib.binding * string) list -> theory -> Proof.state | |
| 48 | val interpret: expression_i -> (Attrib.binding * term) list -> bool -> Proof.state -> Proof.state | |
| 41270 | 49 | val interpret_cmd: expression -> (Attrib.binding * string) list -> | 
| 50 | bool -> Proof.state -> Proof.state | |
| 41435 | 51 | |
| 52 | (* Diagnostic *) | |
| 53 | val print_dependencies: Proof.context -> bool -> expression -> unit | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 54 | end; | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 55 | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 56 | structure Expression : EXPRESSION = | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 57 | struct | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 58 | |
| 28795 | 59 | datatype ctxt = datatype Element.ctxt; | 
| 60 | ||
| 61 | ||
| 62 | (*** Expressions ***) | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 63 | |
| 28872 | 64 | datatype 'term map = | 
| 65 | Positional of 'term option list | | |
| 66 | Named of (string * 'term) list; | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 67 | |
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 68 | type ('name, 'term) expr = ('name * ((string * bool) * 'term map)) list;
 | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 69 | |
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 70 | type expression_i = (string, term) expr * (binding * typ option * mixfix) list; | 
| 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 71 | type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list; | 
| 28795 | 72 | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 73 | |
| 28859 | 74 | (** Internalise locale names in expr **) | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 75 | |
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 76 | 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 | 77 | |
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 78 | |
| 30778 | 79 | (** Parameters of expression **) | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 80 | |
| 30778 | 81 | (*Sanity check of instantiations and extraction of implicit parameters. | 
| 82 | The latter only occurs iff strict = false. | |
| 83 | Positional instantiations are extended to match full length of parameter list | |
| 84 | of instantiated locale.*) | |
| 28895 | 85 | |
| 86 | fun parameters_of thy strict (expr, fixed) = | |
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 87 | let | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 88 | 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 | 89 | (case duplicates (op =) xs of | 
| 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 90 | [] => () | 
| 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 91 | | dups => error (message ^ commas dups)); | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 92 | |
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 93 | fun parm_eq ((p1: string, mx1: mixfix), (p2, mx2)) = p1 = p2 andalso | 
| 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 94 |       (mx1 = mx2 orelse error ("Conflicting syntax for parameter " ^ quote p1 ^ " in expression"));
 | 
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 95 | |
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 96 | fun params_loc loc = Locale.params_of thy loc |> map (apfst #1); | 
| 30778 | 97 | fun params_inst (loc, (prfx, Positional insts)) = | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 98 | let | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 99 | val ps = params_loc loc; | 
| 29358 | 100 | val d = length ps - length insts; | 
| 101 | val insts' = | |
| 102 |               if d < 0 then error ("More arguments than parameters in instantiation of locale " ^
 | |
| 29360 | 103 | quote (Locale.extern thy loc)) | 
| 29358 | 104 | else insts @ replicate d NONE; | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 105 | val ps' = (ps ~~ insts') |> | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 106 | map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE); | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 107 | in (ps', (loc, (prfx, Positional insts'))) end | 
| 30778 | 108 | | params_inst (loc, (prfx, Named insts)) = | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 109 | let | 
| 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 110 | val _ = reject_dups "Duplicate instantiation of the following parameter(s): " | 
| 28859 | 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")); | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 115 | in (ps', (loc, (prfx, Named insts))) 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 | 
| 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 133 | let val _ = reject_dups | 
| 28895 | 134 | "Parameter(s) declared simultaneously in expression and for clause: " (implicit' @ fixed') | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 135 | 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 | 136 | |
| 28895 | 137 | in (expr', implicit'' @ fixed) end; | 
| 28697 
140bfb63f893
New-style locale expressions with instantiation (new file expression.ML).
 ballarin parents: diff
changeset | 138 | |
| 28795 | 139 | |
| 140 | (** Read instantiation **) | |
| 141 | ||
| 28872 | 142 | (* Parse positional or named instantiation *) | 
| 143 | ||
| 28859 | 144 | local | 
| 145 | ||
| 29797 | 146 | fun prep_inst prep_term ctxt parms (Positional insts) = | 
| 49817 | 147 | (insts ~~ parms) |> map | 
| 148 | (fn (NONE, p) => Free (p, dummyT) | |
| 149 | | (SOME t, _) => prep_term ctxt t) | |
| 29797 | 150 | | prep_inst prep_term ctxt parms (Named insts) = | 
| 49817 | 151 | parms |> map (fn p => | 
| 152 | (case AList.lookup (op =) insts p of | |
| 153 | SOME t => prep_term ctxt t | | |
| 154 | NONE => Free (p, dummyT))); | |
| 28872 | 155 | |
| 156 | in | |
| 157 | ||
| 158 | fun parse_inst x = prep_inst Syntax.parse_term x; | |
| 159 | fun make_inst x = prep_inst (K I) x; | |
| 160 | ||
| 161 | end; | |
| 162 | ||
| 163 | ||
| 164 | (* Instantiation morphism *) | |
| 165 | ||
| 30774 | 166 | fun inst_morph (parm_names, parm_types) ((prfx, mandatory), insts') ctxt = | 
| 28795 | 167 | let | 
| 168 | (* parameters *) | |
| 169 | val type_parm_names = fold Term.add_tfreesT parm_types [] |> map fst; | |
| 170 | ||
| 171 | (* type inference and contexts *) | |
| 37145 
01aa36932739
renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
 wenzelm parents: 
36674diff
changeset | 172 | val parm_types' = map (Type_Infer.paramify_vars o Logic.varifyT_global) parm_types; | 
| 28795 | 173 | val type_parms = fold Term.add_tvarsT parm_types' [] |> map (Logic.mk_type o TVar); | 
| 39288 | 174 | val arg = type_parms @ map2 Type.constraint parm_types' insts'; | 
| 28795 | 175 | val res = Syntax.check_terms ctxt arg; | 
| 176 | val ctxt' = ctxt |> fold Variable.auto_fixes res; | |
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 177 | |
| 28795 | 178 | (* instantiation *) | 
| 179 | val (type_parms'', res') = chop (length type_parms) res; | |
| 180 | val insts'' = (parm_names ~~ res') |> map_filter | |
| 30776 | 181 | (fn inst as (x, Free (y, _)) => if x = y then NONE else SOME inst | 
| 182 | | inst => SOME inst); | |
| 28795 | 183 | val instT = Symtab.make (type_parm_names ~~ map Logic.dest_type type_parms''); | 
| 184 | val inst = Symtab.make insts''; | |
| 185 | in | |
| 42360 | 186 | (Element.inst_morphism (Proof_Context.theory_of ctxt) (instT, inst) $> | 
| 30774 | 187 | Morphism.binding_morphism (Binding.prefix mandatory prfx), ctxt') | 
| 28795 | 188 | end; | 
| 28859 | 189 | |
| 28795 | 190 | |
| 191 | (*** Locale processing ***) | |
| 192 | ||
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 193 | (** Parsing **) | 
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 194 | |
| 29604 | 195 | fun parse_elem prep_typ prep_term ctxt = | 
| 196 | Element.map_ctxt | |
| 197 |    {binding = I,
 | |
| 198 | typ = prep_typ ctxt, | |
| 42360 | 199 | term = prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt), | 
| 200 | pattern = prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt), | |
| 29604 | 201 | fact = I, | 
| 202 | attrib = I}; | |
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 203 | |
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 204 | fun parse_concl prep_term ctxt concl = | 
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 205 | (map o map) (fn (t, ps) => | 
| 42360 | 206 | (prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) t, | 
| 207 | map (prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt)) ps)) concl; | |
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 208 | |
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 209 | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 210 | (** Simultaneous type inference: instantiations + elements + conclusion **) | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 211 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 212 | local | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 213 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 214 | fun mk_type T = (Logic.mk_type T, []); | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 215 | fun mk_term t = (t, []); | 
| 39288 | 216 | fun mk_propp (p, pats) = (Type.constraint propT p, pats); | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 217 | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 218 | fun dest_type (T, []) = Logic.dest_type T; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 219 | fun dest_term (t, []) = t; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 220 | fun dest_propp (p, pats) = (p, pats); | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 221 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 222 | fun extract_inst (_, (_, ts)) = map mk_term ts; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 223 | 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 | 224 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 225 | 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 | 226 | | 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 | 227 | | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 228 | | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 229 | | extract_elem (Notes _) = []; | 
| 28795 | 230 | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 231 | fun restore_elem (Fixes fixes, css) = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 232 | (fixes ~~ css) |> map (fn ((x, _, mx), cs) => | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 233 | (x, cs |> map dest_type |> try hd, mx)) |> Fixes | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 234 | | restore_elem (Constrains csts, css) = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 235 | (csts ~~ css) |> map (fn ((x, _), cs) => | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 236 | (x, cs |> map dest_type |> hd)) |> Constrains | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 237 | | restore_elem (Assumes asms, css) = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 238 | (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 | 239 | | restore_elem (Defines defs, css) = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 240 | (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 241 | | restore_elem (Notes notes, _) = Notes notes; | 
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 242 | |
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 243 | fun check cs context = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 244 | let | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 245 | fun prep (_, pats) (ctxt, t :: ts) = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 246 | let val ctxt' = Variable.auto_fixes t ctxt | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 247 | in | 
| 42360 | 248 | ((t, Syntax.check_props (Proof_Context.set_mode Proof_Context.mode_pattern ctxt') pats), | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 249 | (ctxt', ts)) | 
| 30776 | 250 | end; | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 251 | val (cs', (context', _)) = fold_map prep cs | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 252 | (context, Syntax.check_terms | 
| 42360 | 253 | (Proof_Context.set_mode Proof_Context.mode_schematic context) (map fst cs)); | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 254 | in (cs', context') end; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 255 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 256 | in | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 257 | |
| 28872 | 258 | fun check_autofix insts elems concl ctxt = | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 259 | let | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 260 | val inst_cs = map extract_inst insts; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 261 | val elem_css = map extract_elem elems; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 262 | val concl_cs = (map o map) mk_propp concl; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 263 | (* Type inference *) | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 264 | val (inst_cs' :: css', ctxt') = | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 265 | (fold_burrow o fold_burrow) check (inst_cs :: elem_css @ [concl_cs]) ctxt; | 
| 28936 
f1647bf418f5
No resolution of patterns within context statements.
 ballarin parents: 
28903diff
changeset | 266 | 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 | 267 | in | 
| 30776 | 268 | (map restore_inst (insts ~~ inst_cs'), | 
| 269 | map restore_elem (elems ~~ elem_css'), | |
| 28936 
f1647bf418f5
No resolution of patterns within context statements.
 ballarin parents: 
28903diff
changeset | 270 | concl_cs', ctxt') | 
| 28885 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 271 | end; | 
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 272 | |
| 
6f6bf52e75bb
Expression types cleaned up, proper treatment of term patterns.
 ballarin parents: 
28879diff
changeset | 273 | end; | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 274 | |
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 275 | |
| 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 276 | (** Prepare locale elements **) | 
| 28795 | 277 | |
| 278 | fun declare_elem prep_vars (Fixes fixes) ctxt = | |
| 279 | let val (vars, _) = prep_vars fixes ctxt | |
| 42360 | 280 | in ctxt |> Proof_Context.add_fixes vars |> snd end | 
| 28795 | 281 | | declare_elem prep_vars (Constrains csts) ctxt = | 
| 28965 | 282 | ctxt |> prep_vars (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) csts) |> snd | 
| 28872 | 283 | | declare_elem _ (Assumes _) ctxt = ctxt | 
| 284 | | declare_elem _ (Defines _) ctxt = ctxt | |
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 285 | | declare_elem _ (Notes _) ctxt = ctxt; | 
| 28795 | 286 | |
| 30776 | 287 | |
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 288 | (** Finish locale elements **) | 
| 28795 | 289 | |
| 49819 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 290 | 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 | 291 | let | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 292 | val thy = Proof_Context.theory_of ctxt; | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 293 | val (parm_names, parm_types) = Locale.params_of thy loc |> map #1 |> split_list; | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 294 | val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst) ctxt; | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 295 | 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 | 296 | |
| 49818 | 297 | fun finish_fixes (parms: (string * typ) list) = map (fn (binding, _, mx) => | 
| 298 | let val x = Binding.name_of binding | |
| 299 | in (binding, AList.lookup (op =) parms x, mx) end); | |
| 49817 | 300 | |
| 301 | local | |
| 302 | ||
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 303 | 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 | 304 | | closeup (outer_ctxt, ctxt) parms true elem = | 
| 28795 | 305 | let | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 306 | (* FIXME consider closing in syntactic phase -- before type checking *) | 
| 28795 | 307 | fun close_frees t = | 
| 308 | let | |
| 309 | val rev_frees = | |
| 310 | 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 | 311 | 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 | 312 | else insert (op =) (x, T) | _ => I) t []; | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 313 | in fold (Logic.all o Free) rev_frees t end; | 
| 28795 | 314 | |
| 315 | fun no_binds [] = [] | |
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 316 | | no_binds _ = error "Illegal term bindings in context element"; | 
| 28795 | 317 | in | 
| 318 | (case elem of | |
| 319 | Assumes asms => Assumes (asms |> map (fn (a, propps) => | |
| 320 | (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps))) | |
| 29022 | 321 | | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) => | 
| 35624 | 322 | let val ((c, _), t') = Local_Defs.cert_def ctxt (close_frees t) | 
| 30434 | 323 | in ((Thm.def_binding_optional (Binding.name c) name, atts), (t', no_binds ps)) end)) | 
| 28795 | 324 | | e => e) | 
| 325 | end; | |
| 326 | ||
| 49819 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 327 | in | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 328 | |
| 49818 | 329 | fun finish_elem _ parms _ (Fixes fixes) = Fixes (finish_fixes parms fixes) | 
| 330 | | 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 | 331 | | 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 | 332 | | finish_elem ctxts parms do_close (Defines defs) = closeup ctxts parms do_close (Defines defs) | 
| 49818 | 333 | | finish_elem _ _ _ (Notes facts) = Notes facts; | 
| 28872 | 334 | |
| 49817 | 335 | end; | 
| 336 | ||
| 28795 | 337 | |
| 28895 | 338 | (** Process full context statement: instantiations + elements + conclusion **) | 
| 339 | ||
| 340 | (* Interleave incremental parsing and type inference over entire parsed stretch. *) | |
| 341 | ||
| 28795 | 342 | local | 
| 343 | ||
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 344 | fun prep_full_context_statement | 
| 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 345 | parse_typ parse_prop prep_vars_elem prep_inst prep_vars_inst prep_expr | 
| 30786 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 346 |     {strict, do_close, fixed_frees} raw_import init_body raw_elems raw_concl ctxt1 =
 | 
| 28795 | 347 | let | 
| 42360 | 348 | val thy = Proof_Context.theory_of ctxt1; | 
| 28872 | 349 | |
| 28895 | 350 | val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import); | 
| 351 | ||
| 30778 | 352 | fun prep_insts_cumulative (loc, (prfx, inst)) (i, insts, ctxt) = | 
| 28872 | 353 | let | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 354 | val (parm_names, parm_types) = Locale.params_of thy loc |> map #1 |> split_list; | 
| 29797 | 355 | val inst' = prep_inst ctxt parm_names inst; | 
| 45587 | 356 | val parm_types' = parm_types | 
| 357 | |> map (Type_Infer.paramify_vars o | |
| 358 | Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) o Logic.varifyT_global); | |
| 39288 | 359 | val inst'' = map2 Type.constraint parm_types' inst'; | 
| 28872 | 360 | val insts' = insts @ [(loc, (prfx, inst''))]; | 
| 47315 | 361 | val (insts'', _, _, _) = check_autofix insts' [] [] ctxt; | 
| 28872 | 362 | val inst''' = insts'' |> List.last |> snd |> snd; | 
| 363 | val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst''') ctxt; | |
| 30764 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 wenzelm parents: 
30763diff
changeset | 364 | val ctxt'' = Locale.activate_declarations (loc, morph) ctxt; | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 365 | in (i + 1, insts', ctxt'') end; | 
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 366 | |
| 47315 | 367 | fun prep_elem insts raw_elem ctxt = | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 368 | let | 
| 47315 | 369 | val ctxt' = ctxt | 
| 370 | |> Context_Position.set_visible false | |
| 371 | |> declare_elem prep_vars_elem raw_elem | |
| 372 | |> Context_Position.restore_visible ctxt; | |
| 373 | 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 | 374 | in (elems', ctxt') end; | 
| 28795 | 375 | |
| 28872 | 376 | fun prep_concl raw_concl (insts, elems, ctxt) = | 
| 28795 | 377 | let | 
| 29215 
f98862eb0591
Use correct mode when parsing elements and conclusion.
 ballarin parents: 
29214diff
changeset | 378 | val concl = parse_concl parse_prop ctxt raw_concl; | 
| 28872 | 379 | in check_autofix insts elems concl ctxt end; | 
| 28795 | 380 | |
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 381 | val fors = prep_vars_inst fixed ctxt1 |> fst; | 
| 42360 | 382 | val ctxt2 = ctxt1 |> Proof_Context.add_fixes fors |> snd; | 
| 30778 | 383 | val (_, insts', ctxt3) = fold prep_insts_cumulative raw_insts (0, [], ctxt2); | 
| 30786 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 384 | |
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 385 | val _ = | 
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 386 | if fixed_frees then () | 
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 387 | else | 
| 42482 | 388 | (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 | 389 | [] => () | 
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 390 |         | frees => error ("Illegal free variables in expression: " ^
 | 
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 391 | 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 | 392 | |
| 29702 | 393 | val ctxt4 = init_body ctxt3; | 
| 47315 | 394 | val (elems, ctxt5) = fold_map (prep_elem insts') raw_elems ctxt4; | 
| 30778 | 395 | val (insts, elems', concl, ctxt6) = prep_concl raw_concl (insts', elems, ctxt5); | 
| 28795 | 396 | |
| 28872 | 397 | (* Retrieve parameter types *) | 
| 42494 | 398 | val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Variable.check_name o #1) fixes) | 
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 399 | | _ => fn ps => ps) (Fixes fors :: elems') []; | 
| 42360 | 400 | val (Ts, ctxt7) = fold_map Proof_Context.inferred_param xs ctxt6; | 
| 28895 | 401 | val parms = xs ~~ Ts; (* params from expression and elements *) | 
| 28795 | 402 | |
| 49818 | 403 | 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 | 404 | val fixed = map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fors'; | 
| 49819 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 405 | val deps = map (finish_inst ctxt6) insts; | 
| 
97b572c10fe9
refrain from quantifying outer fixes, to enable nesting of contexts like "context fixes x context assumes A x";
 wenzelm parents: 
49818diff
changeset | 406 | val elems'' = map (finish_elem (ctxt1, ctxt6) parms do_close) elems'; | 
| 28852 
5ddea758679b
Type inference for elements through syntax module.
 ballarin parents: 
28832diff
changeset | 407 | |
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 408 | in ((fixed, deps, elems'', concl), (parms, ctxt7)) end | 
| 28795 | 409 | |
| 410 | in | |
| 411 | ||
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 412 | fun cert_full_context_statement x = | 
| 42360 | 413 | prep_full_context_statement (K I) (K I) Proof_Context.cert_vars | 
| 414 | make_inst Proof_Context.cert_vars (K I) x; | |
| 30776 | 415 | |
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 416 | fun cert_read_full_context_statement x = | 
| 42360 | 417 | prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Proof_Context.read_vars | 
| 418 | make_inst Proof_Context.cert_vars (K I) x; | |
| 30776 | 419 | |
| 28895 | 420 | fun read_full_context_statement x = | 
| 42360 | 421 | prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Proof_Context.read_vars | 
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 422 | parse_inst Proof_Context.read_vars check_expr x; | 
| 28795 | 423 | |
| 424 | end; | |
| 425 | ||
| 426 | ||
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 427 | (* Context statement: elements + conclusion *) | 
| 28795 | 428 | |
| 429 | local | |
| 430 | ||
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 431 | fun prep_statement prep activate raw_elems raw_concl context = | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 432 | let | 
| 49817 | 433 | val ((_, _, elems, concl), _) = | 
| 434 |       prep {strict = true, do_close = false, fixed_frees = true}
 | |
| 30786 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 435 | ([], []) I raw_elems raw_concl context; | 
| 49817 | 436 | val (_, context') = context | 
| 437 | |> Proof_Context.set_stmt true | |
| 438 | |> fold_map activate elems; | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 439 | in (concl, context') end; | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 440 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 441 | in | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 442 | |
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 443 | 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 | 444 | 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 | 445 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 446 | end; | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 447 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 448 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 449 | (* Locale declaration: import + elements *) | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 450 | |
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 451 | fun fix_params params = | 
| 42360 | 452 | 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 | 453 | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 454 | local | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 455 | |
| 29702 | 456 | fun prep_declaration prep activate raw_import init_body raw_elems context = | 
| 28795 | 457 | let | 
| 29358 | 458 | val ((fixed, deps, elems, _), (parms, ctxt')) = | 
| 30786 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 459 |       prep {strict = false, do_close = true, fixed_frees = false}
 | 
| 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 460 | raw_import init_body raw_elems [] context; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 461 | (* Declare parameters and imported facts *) | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 462 | val context' = context |> | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 463 | fix_params fixed |> | 
| 38316 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 ballarin parents: 
38211diff
changeset | 464 | fold (Context.proof_map o Locale.activate_facts NONE) deps; | 
| 47311 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 wenzelm parents: 
47287diff
changeset | 465 | val (elems', context'') = context' |> | 
| 42360 | 466 | Proof_Context.set_stmt true |> | 
| 30777 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30776diff
changeset | 467 | fold_map activate elems; | 
| 47311 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 wenzelm parents: 
47287diff
changeset | 468 | in ((fixed, deps, elems', context''), (parms, ctxt')) end; | 
| 28795 | 469 | |
| 470 | in | |
| 471 | ||
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 472 | 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 | 473 | 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 | 474 | 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 | 475 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 476 | end; | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 477 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 478 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 479 | (* Locale expression to set up a goal *) | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 480 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 481 | local | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 482 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 483 | fun props_of thy (name, morph) = | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 484 | let | 
| 29360 | 485 | val (asm, defs) = Locale.specification_of thy name; | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 486 | in | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 487 | (case asm of NONE => defs | SOME asm => asm :: defs) |> map (Morphism.term morph) | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 488 | end; | 
| 
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 | fun prep_goal_expression prep expression context = | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 491 | let | 
| 42360 | 492 | val thy = Proof_Context.theory_of context; | 
| 28879 | 493 | |
| 29358 | 494 | val ((fixed, deps, _, _), _) = | 
| 30786 
461f7b5f16a2
prep_full_context_statement: explicit record of flags;
 wenzelm parents: 
30784diff
changeset | 495 |       prep {strict = true, do_close = true, fixed_frees = true} expression I [] [] context;
 | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 496 | (* proof obligations *) | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 497 | val propss = map (props_of thy) deps; | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 498 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 499 | val goal_ctxt = context |> | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 500 | fix_params fixed |> | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 501 | (fold o fold) Variable.auto_fixes propss; | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 502 | |
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 503 | val export = Variable.export_morphism goal_ctxt context; | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 504 | val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export; | 
| 31977 | 505 | 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 | 506 | val exp_typ = Logic.type_map exp_term; | 
| 45289 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 wenzelm parents: 
43543diff
changeset | 507 | val export' = | 
| 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 wenzelm parents: 
43543diff
changeset | 508 |       Morphism.morphism {binding = [], typ = [exp_typ], term = [exp_term], fact = [exp_fact]};
 | 
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 509 | in ((propss, deps, export'), goal_ctxt) end; | 
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 510 | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 511 | in | 
| 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 512 | |
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 513 | 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 | 514 | fun read_goal_expression x = prep_goal_expression read_full_context_statement x; | 
| 28879 | 515 | |
| 28795 | 516 | end; | 
| 517 | ||
| 518 | ||
| 519 | (*** Locale declarations ***) | |
| 520 | ||
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 521 | (* extract specification text *) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 522 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 523 | val norm_term = Envir.beta_norm oo Term.subst_atomic; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 524 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 525 | fun bind_def ctxt eq (xs, env, eqs) = | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 526 | let | 
| 35624 | 527 | val _ = Local_Defs.cert_def ctxt eq; | 
| 528 | val ((y, T), b) = Local_Defs.abs_def eq; | |
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 529 | val b' = norm_term env b; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 530 | fun err msg = error (msg ^ ": " ^ quote y); | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 531 | in | 
| 49749 | 532 | (case filter (fn (Free (y', _), _) => y = y' | _ => false) env of | 
| 533 | [] => (Term.add_frees b' xs, (Free (y, T), b') :: env, eq :: eqs) | |
| 534 | | dups => | |
| 535 | if forall (fn (_, b'') => b' aconv b'') dups then (xs, env, eqs) | |
| 536 | else err "Attempt to redefine variable") | |
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 537 | end; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 538 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 539 | (* text has the following structure: | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 540 | (((exts, exts'), (ints, ints')), (xs, env, defs)) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 541 | where | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 542 | exts: external assumptions (terms in assumes elements) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 543 | exts': dito, normalised wrt. env | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 544 | ints: internal assumptions (terms in assumptions from insts) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 545 | ints': dito, normalised wrt. env | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 546 | 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 | 547 | this includes parameters except defined parameters | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 548 | env: list of term pairs encoding substitutions, where the first term | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 549 | is a free variable; substitutions represent defines elements and | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 550 | the rhs is normalised wrt. the previous env | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 551 | defs: the equations from the defines elements | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 552 | *) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 553 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 554 | fun eval_text _ _ (Fixes _) text = text | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 555 | | eval_text _ _ (Constrains _) text = text | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 556 | | eval_text _ is_ext (Assumes asms) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 557 | (((exts, exts'), (ints, ints')), (xs, env, defs)) = | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 558 | let | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 559 | val ts = maps (map #1 o #2) asms; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 560 | val ts' = map (norm_term env) ts; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 561 | val spec' = | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 562 | if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints')) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 563 | else ((exts, exts'), (ints @ ts, ints' @ ts')); | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 564 | in (spec', (fold Term.add_frees ts' xs, env, defs)) end | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 565 | | eval_text ctxt _ (Defines defs) (spec, binds) = | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 566 | (spec, fold (bind_def ctxt o #1 o #2) defs binds) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 567 | | eval_text _ _ (Notes _) text = text; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 568 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 569 | fun eval_inst ctxt (loc, morph) text = | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 570 | let | 
| 42360 | 571 | val thy = Proof_Context.theory_of ctxt; | 
| 29360 | 572 | val (asm, defs) = Locale.specification_of thy loc; | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 573 | val asm' = Option.map (Morphism.term morph) asm; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 574 | val defs' = map (Morphism.term morph) defs; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 575 | val text' = text |> | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 576 | (if is_some asm | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 577 | then eval_text ctxt false (Assumes [(Attrib.empty_binding, [(the asm', [])])]) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 578 | else I) |> | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 579 | (if not (null defs) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 580 | then eval_text ctxt false (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs')) | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 581 | else I) | 
| 29360 | 582 | (* FIXME clone from locale.ML *) | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 583 | in text' end; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 584 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 585 | fun eval_elem ctxt elem text = | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 586 | eval_text ctxt true elem text; | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 587 | |
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 588 | fun eval ctxt deps elems = | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 589 | let | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 590 | val text' = fold (eval_inst ctxt) deps ((([], []), ([], [])), ([], [], [])); | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 591 | val ((spec, (_, _, defs))) = fold (eval_elem ctxt) elems text'; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 592 | in (spec, defs) end; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 593 | |
| 28903 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 594 | (* 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 | 595 | also name suffix of delta predicates and assumptions. *) | 
| 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 596 | |
| 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 597 | val axiomsN = "axioms"; | 
| 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 598 | |
| 28795 | 599 | local | 
| 600 | ||
| 601 | (* 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 | 602 | delta predicates *) | 
| 28795 | 603 | |
| 604 | val introN = "intro"; | |
| 605 | ||
| 606 | fun atomize_spec thy ts = | |
| 607 | let | |
| 608 | val t = Logic.mk_conjunction_balanced ts; | |
| 35625 | 609 | val body = Object_Logic.atomize_term thy t; | 
| 28795 | 610 | val bodyT = Term.fastype_of body; | 
| 611 | in | |
| 612 | if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t)) | |
| 35625 | 613 | else (body, bodyT, Object_Logic.atomize (Thm.cterm_of thy t)) | 
| 28795 | 614 | end; | 
| 615 | ||
| 616 | (* achieve plain syntax for locale predicates (without "PROP") *) | |
| 617 | ||
| 49820 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 618 | fun aprop_tr' n c = | 
| 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 619 | let | 
| 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 620 | val c' = Lexicon.mark_const c; | 
| 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 621 | fun tr' T args = | 
| 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 622 | 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 | 623 | 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 | 624 | else raise Match; | 
| 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 625 | in (c', tr') end; | 
| 28795 | 626 | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 627 | (* define one predicate including its intro rule and axioms | 
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 628 | - binding: predicate name | 
| 28795 | 629 | - parms: locale parameters | 
| 630 | - defs: thms representing substitutions from defines elements | |
| 631 | - ts: terms representing locale assumptions (not normalised wrt. defs) | |
| 632 | - norm_ts: terms representing locale assumptions (normalised wrt. defs) | |
| 633 | - thy: the theory | |
| 634 | *) | |
| 635 | ||
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 636 | fun def_pred binding parms defs ts norm_ts thy = | 
| 28795 | 637 | let | 
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 638 | val name = Sign.full_name thy binding; | 
| 28795 | 639 | |
| 640 | val (body, bodyT, body_eq) = atomize_spec thy norm_ts; | |
| 29272 | 641 | val env = Term.add_free_names body []; | 
| 28795 | 642 | val xs = filter (member (op =) env o #1) parms; | 
| 643 | val Ts = map #2 xs; | |
| 29272 | 644 | val extraTs = | 
| 33040 | 645 | (subtract (op =) (fold Term.add_tfreesT Ts []) (Term.add_tfrees body [])) | 
| 28795 | 646 | |> Library.sort_wrt #1 |> map TFree; | 
| 647 | val predT = map Term.itselfT extraTs ---> Ts ---> bodyT; | |
| 648 | ||
| 649 | val args = map Logic.mk_type extraTs @ map Free xs; | |
| 650 | val head = Term.list_comb (Const (name, predT), args); | |
| 35625 | 651 | val statement = Object_Logic.ensure_propT thy head; | 
| 28795 | 652 | |
| 653 | val ([pred_def], defs_thy) = | |
| 654 | thy | |
| 49820 
f7a1e1745b7b
refined aprop_tr' -- retain entity information by using type slot as adhoc marker;
 wenzelm parents: 
49819diff
changeset | 655 | |> bodyT = propT ? Sign.add_trfunsT [aprop_tr' (length args) name] | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42360diff
changeset | 656 | |> Sign.declare_const_global ((Binding.conceal binding, predT), NoSyn) |> snd | 
| 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 | 657 | |> Global_Theory.add_defs false | 
| 35238 | 658 | [((Binding.conceal (Thm.def_binding binding), Logic.mk_equals (head, body)), [])]; | 
| 42360 | 659 | val defs_ctxt = Proof_Context.init_global defs_thy |> Variable.declare_term head; | 
| 28795 | 660 | |
| 661 | val cert = Thm.cterm_of defs_thy; | |
| 662 | ||
| 663 | val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ => | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
39557diff
changeset | 664 | rewrite_goals_tac [pred_def] THEN | 
| 28795 | 665 | Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN | 
| 666 | Tactic.compose_tac (false, | |
| 667 | Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1); | |
| 668 | ||
| 669 | val conjuncts = | |
| 670 | (Drule.equal_elim_rule2 OF [body_eq, | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
39557diff
changeset | 671 | Raw_Simplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))]) | 
| 28795 | 672 | |> Conjunction.elim_balanced (length ts); | 
| 673 | val axioms = ts ~~ conjuncts |> map (fn (t, ax) => | |
| 674 | Element.prove_witness defs_ctxt t | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
39557diff
changeset | 675 | (rewrite_goals_tac defs THEN | 
| 28795 | 676 | Tactic.compose_tac (false, ax, 0) 1)); | 
| 677 | in ((statement, intro, axioms), defs_thy) end; | |
| 678 | ||
| 679 | in | |
| 680 | ||
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30335diff
changeset | 681 | (* main predicate definition function *) | 
| 28795 | 682 | |
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 683 | fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy = | 
| 28795 | 684 | let | 
| 30762 
cabf9ff3a129
define_prefs: removed redundant Drule.gen_all, which is already part of the norm_hhf stage of Assumption.assume;
 wenzelm parents: 
30755diff
changeset | 685 | val defs' = map (cterm_of thy #> Assumption.assume #> Drule.abs_def) defs; | 
| 29031 
e74341997a48
Pass on defines in inheritance; reject illicit defines created by instantiation.
 ballarin parents: 
29030diff
changeset | 686 | |
| 28795 | 687 | val (a_pred, a_intro, a_axioms, thy'') = | 
| 688 | if null exts then (NONE, NONE, [], thy) | |
| 689 | else | |
| 690 | let | |
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 691 |           val abinding = if null ints then binding else Binding.suffix_name ("_" ^ axiomsN) binding;
 | 
| 28795 | 692 | val ((statement, intro, axioms), thy') = | 
| 693 | thy | |
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 694 | |> def_pred abinding parms defs' exts exts'; | 
| 28795 | 695 | val (_, thy'') = | 
| 696 | thy' | |
| 35204 
214ec053128e
locale: more precise treatment of naming vs. binding;
 wenzelm parents: 
35143diff
changeset | 697 | |> Sign.qualified_path true abinding | 
| 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 | 698 | |> Global_Theory.note_thmss "" | 
| 33278 | 699 | [((Binding.conceal (Binding.name introN), []), [([intro], [Locale.unfold_add])])] | 
| 28795 | 700 | ||> Sign.restore_naming thy'; | 
| 701 | in (SOME statement, SOME intro, axioms, thy'') end; | |
| 702 | val (b_pred, b_intro, b_axioms, thy'''') = | |
| 703 | if null ints then (NONE, NONE, [], thy'') | |
| 704 | else | |
| 705 | let | |
| 706 | val ((statement, intro, axioms), thy''') = | |
| 707 | thy'' | |
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 708 | |> def_pred binding parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred); | 
| 28795 | 709 | val (_, thy'''') = | 
| 710 | thy''' | |
| 35204 
214ec053128e
locale: more precise treatment of naming vs. binding;
 wenzelm parents: 
35143diff
changeset | 711 | |> 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 | 712 | |> Global_Theory.note_thmss "" | 
| 33278 | 713 | [((Binding.conceal (Binding.name introN), []), [([intro], [Locale.intro_add])]), | 
| 714 | ((Binding.conceal (Binding.name axiomsN), []), | |
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
33671diff
changeset | 715 | [(map (Drule.export_without_context o Element.conclude_witness) axioms, [])])] | 
| 28795 | 716 | ||> Sign.restore_naming thy'''; | 
| 717 | in (SOME statement, SOME intro, axioms, thy'''') end; | |
| 718 | in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy'''') end; | |
| 719 | ||
| 720 | end; | |
| 721 | ||
| 722 | ||
| 723 | local | |
| 724 | ||
| 725 | fun assumes_to_notes (Assumes asms) axms = | |
| 726 | fold_map (fn (a, spec) => fn axs => | |
| 727 | let val (ps, qs) = chop (length spec) axs | |
| 728 | in ((a, [(ps, [])]), qs) end) asms axms | |
| 33644 
5266a72e0889
eliminated slightly odd (unused) "axiom" and "assumption" -- collapsed to unspecific "";
 wenzelm parents: 
33643diff
changeset | 729 | |> apfst (curry Notes "") | 
| 28795 | 730 | | assumes_to_notes e axms = (e, axms); | 
| 731 | ||
| 29031 
e74341997a48
Pass on defines in inheritance; reject illicit defines created by instantiation.
 ballarin parents: 
29030diff
changeset | 732 | fun defines_to_notes thy (Defines defs) = | 
| 42440 
5e7a7343ab11
discontinuend obsolete Thm.definitionK, which was hardly ever well-defined;
 wenzelm parents: 
42381diff
changeset | 733 |       Notes ("", map (fn (a, (def, _)) =>
 | 
| 29245 
19728ee2b1ba
Intro_locales_tac knows about defines elements; more robust export morphism.
 ballarin parents: 
29241diff
changeset | 734 | (a, [([Assumption.assume (cterm_of thy def)], | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 735 | [(Attrib.internal o K) Locale.witness_add])])) defs) | 
| 29031 
e74341997a48
Pass on defines in inheritance; reject illicit defines created by instantiation.
 ballarin parents: 
29030diff
changeset | 736 | | defines_to_notes _ e = e; | 
| 28795 | 737 | |
| 28898 
530c7d28a962
Proper treatment of expressions with free arguments.
 ballarin parents: 
28895diff
changeset | 738 | fun gen_add_locale prep_decl | 
| 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 | 739 | before_exit binding raw_predicate_binding raw_import raw_body thy = | 
| 28795 | 740 | let | 
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 741 | val name = Sign.full_name thy binding; | 
| 29391 
1f6e8c00dc3e
tuned signature; changed locale predicate name convention
 haftmann parents: 
29362diff
changeset | 742 | val _ = Locale.defined thy name andalso | 
| 28795 | 743 |       error ("Duplicate definition of locale " ^ quote name);
 | 
| 744 | ||
| 47311 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 wenzelm parents: 
47287diff
changeset | 745 | val ((fixed, deps, body_elems, _), (parms, ctxt')) = | 
| 42360 | 746 | prep_decl raw_import I raw_body (Proof_Context.init_global thy); | 
| 29221 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 747 | val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems; | 
| 
918687637307
Refactored: evaluate specification text only in locale declarations.
 ballarin parents: 
29217diff
changeset | 748 | |
| 37313 | 749 | val extraTs = | 
| 750 | subtract (op =) (fold Term.add_tfreesT (map snd parms) []) (fold Term.add_tfrees exts' []); | |
| 751 | val _ = | |
| 752 | if null extraTs then () | |
| 753 |       else warning ("Additional type variable(s) in locale specification " ^
 | |
| 42381 
309ec68442c6
added Binding.print convenience, which includes quote already;
 wenzelm parents: 
42375diff
changeset | 754 | Binding.print binding ^ ": " ^ | 
| 37313 | 755 | commas (map (Syntax.string_of_typ ctxt' o TFree) (sort_wrt #1 extraTs))); | 
| 756 | ||
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 757 | val predicate_binding = | 
| 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 758 | if Binding.is_empty raw_predicate_binding then binding | 
| 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 759 | else raw_predicate_binding; | 
| 28872 | 760 | 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 | 761 | define_preds predicate_binding parms text thy; | 
| 28795 | 762 | |
| 29035 | 763 | val a_satisfy = Element.satisfy_morphism a_axioms; | 
| 764 | val b_satisfy = Element.satisfy_morphism b_axioms; | |
| 28903 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 765 | |
| 28895 | 766 | val params = fixed @ | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 767 | maps (fn Fixes fixes => | 
| 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30725diff
changeset | 768 | 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 | 769 | 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 | 770 | |
| 
b5dad96c755a
When adding locales, delay notes until local theory is built.
 ballarin parents: 
29022diff
changeset | 771 | val notes = | 
| 33278 | 772 | if is_some asm then | 
| 33643 
b275f26a638b
eliminated obsolete "internal" kind -- collapsed to unspecific "";
 wenzelm parents: 
33553diff
changeset | 773 |         [("", [((Binding.conceal (Binding.suffix_name ("_" ^ axiomsN) binding), []),
 | 
| 33278 | 774 | [([Assumption.assume (cterm_of thy' (the asm))], | 
| 775 | [(Attrib.internal o K) Locale.witness_add])])])] | |
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 776 | else []; | 
| 28795 | 777 | |
| 29035 | 778 | val notes' = body_elems |> | 
| 779 | map (defines_to_notes thy') |> | |
| 45290 | 780 | map (Element.transform_ctxt a_satisfy) |> | 
| 29035 | 781 | (fn elems => fold_map assumes_to_notes elems (map Element.conclude_witness a_axioms)) |> | 
| 782 | fst |> | |
| 45290 | 783 | map (Element.transform_ctxt b_satisfy) |> | 
| 29035 | 784 | map_filter (fn Notes notes => SOME notes | _ => NONE); | 
| 785 | ||
| 786 | val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps; | |
| 29441 | 787 | val axioms = map Element.conclude_witness b_axioms; | 
| 28872 | 788 | |
| 29358 | 789 | val loc_ctxt = thy' | 
| 33360 
f7d9c5e5d2f9
tuned variable names of bindings; conceal predicate constants
 haftmann parents: 
33315diff
changeset | 790 | |> Locale.register_locale binding (extraTs, params) | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35625diff
changeset | 791 | (asm, rev defs) (a_intro, b_intro) axioms [] (rev notes) (rev deps') | 
| 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 | 792 | |> Named_Target.init before_exit name | 
| 33671 | 793 | |> 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 | 794 | |
| 29358 | 795 | in (name, loc_ctxt) end; | 
| 28795 | 796 | |
| 797 | in | |
| 798 | ||
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 799 | val add_locale = gen_add_locale cert_declaration; | 
| 28902 
2019bcc9d8bf
Ahere to modern naming conventions; proper treatment of internal vs external names.
 ballarin parents: 
28898diff
changeset | 800 | val add_locale_cmd = gen_add_locale read_declaration; | 
| 28795 | 801 | |
| 802 | end; | |
| 803 | ||
| 28895 | 804 | |
| 805 | (*** Interpretation ***) | |
| 806 | ||
| 43543 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 807 | fun read_with_extended_syntax parse_prop deps ctxt props = | 
| 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 808 | let | 
| 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 809 | val deps_ctxt = fold Locale.activate_declarations deps ctxt; | 
| 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 810 | in | 
| 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 811 | map (parse_prop deps_ctxt o snd) props |> Syntax.check_terms deps_ctxt | 
| 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 812 | |> Variable.export_terms deps_ctxt ctxt | 
| 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 813 | end; | 
| 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 814 | |
| 45588 | 815 | fun meta_rewrite ctxt = map (Local_Defs.meta_rewrite_rule ctxt #> Drule.abs_def); | 
| 816 | ||
| 817 | ||
| 38108 | 818 | (** Interpretation in theories and proof contexts **) | 
| 28993 
829e684b02ef
Interpretation in theories including interaction with subclass relation.
 ballarin parents: 
28951diff
changeset | 819 | |
| 
829e684b02ef
Interpretation in theories including interaction with subclass relation.
 ballarin parents: 
28951diff
changeset | 820 | local | 
| 
829e684b02ef
Interpretation in theories including interaction with subclass relation.
 ballarin parents: 
28951diff
changeset | 821 | |
| 46858 | 822 | fun note_eqns_register deps witss attrss eqns export export' = | 
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47068diff
changeset | 823 | Attrib.generic_notes Thm.lemmaK | 
| 45588 | 824 | (attrss ~~ map (fn eqn => [([Morphism.thm (export' $> export) eqn], [])]) eqns) | 
| 46858 | 825 | #-> (fn facts => `(fn context => meta_rewrite (Context.proof_of context) (maps snd facts))) | 
| 826 | #-> (fn eqns => fold (fn ((dep, morph), wits) => | |
| 45588 | 827 | fn context => | 
| 46858 | 828 | Locale.add_registration | 
| 829 | (dep, morph $> Element.satisfy_morphism (map (Element.transform_witness export') wits)) | |
| 830 | (Element.eq_morphism (Context.theory_of context) eqns |> Option.map (rpair true)) | |
| 45588 | 831 | export context) (deps ~~ witss)); | 
| 38108 | 832 | |
| 29211 | 833 | fun gen_interpretation prep_expr parse_prop prep_attr | 
| 45588 | 834 | expression equations thy = | 
| 28993 
829e684b02ef
Interpretation in theories including interaction with subclass relation.
 ballarin parents: 
28951diff
changeset | 835 | let | 
| 45588 | 836 | val ((propss, deps, export), expr_ctxt) = Proof_Context.init_global thy | 
| 29496 | 837 | |> prep_expr expression; | 
| 43543 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 838 | val eqns = read_with_extended_syntax parse_prop deps expr_ctxt equations; | 
| 29441 | 839 | |
| 45588 | 840 | val attrss = map (apsnd (map (prep_attr thy)) o fst) equations; | 
| 29211 | 841 | val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt; | 
| 842 | val export' = Variable.export_morphism goal_ctxt expr_ctxt; | |
| 29210 
4025459e3f83
Interpretation in theories: first version with equations.
 ballarin parents: 
29208diff
changeset | 843 | |
| 38757 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 wenzelm parents: 
38756diff
changeset | 844 | fun after_qed witss eqns = | 
| 42360 | 845 | (Proof_Context.background_theory o Context.theory_map) | 
| 38757 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 wenzelm parents: 
38756diff
changeset | 846 | (note_eqns_register deps witss attrss eqns export export'); | 
| 29441 | 847 | |
| 29578 | 848 | in Element.witness_proof_eqs after_qed propss eqns goal_ctxt end; | 
| 28993 
829e684b02ef
Interpretation in theories including interaction with subclass relation.
 ballarin parents: 
28951diff
changeset | 849 | |
| 38108 | 850 | fun gen_interpret prep_expr parse_prop prep_attr | 
| 851 | expression equations int state = | |
| 852 | let | |
| 853 | val _ = Proof.assert_forward_or_chain state; | |
| 854 | val ctxt = Proof.context_of state; | |
| 45588 | 855 | val thy = Proof_Context.theory_of ctxt; | 
| 38108 | 856 | |
| 857 | val ((propss, deps, export), expr_ctxt) = prep_expr expression ctxt; | |
| 43543 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 858 | val eqns = read_with_extended_syntax parse_prop deps expr_ctxt equations; | 
| 38108 | 859 | |
| 45588 | 860 | val attrss = map (apsnd (map (prep_attr thy)) o fst) equations; | 
| 38108 | 861 | val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt; | 
| 862 | val export' = Variable.export_morphism goal_ctxt expr_ctxt; | |
| 863 | ||
| 45588 | 864 | fun after_qed witss eqns = | 
| 865 | (Proof.map_context o Context.proof_map) | |
| 46899 
58110c1e02bc
refined 'interpret': reset facts ("this") and print_result, which merely consist of internal / protected statement;
 wenzelm parents: 
46858diff
changeset | 866 | (note_eqns_register deps witss attrss eqns export export') | 
| 47068 | 867 | #> Proof.reset_facts; | 
| 38108 | 868 | in | 
| 869 | state | |
| 870 | |> Element.witness_local_proof_eqs after_qed "interpret" propss eqns goal_ctxt int | |
| 871 | end; | |
| 872 | ||
| 28993 
829e684b02ef
Interpretation in theories including interaction with subclass relation.
 ballarin parents: 
28951diff
changeset | 873 | in | 
| 
829e684b02ef
Interpretation in theories including interaction with subclass relation.
 ballarin parents: 
28951diff
changeset | 874 | |
| 29501 
08df2e51cb80
added cert_read_declaration; more exports; tuned signature
 haftmann parents: 
29496diff
changeset | 875 | fun interpretation x = gen_interpretation cert_goal_expression (K I) (K I) x; | 
| 29210 
4025459e3f83
Interpretation in theories: first version with equations.
 ballarin parents: 
29208diff
changeset | 876 | fun interpretation_cmd x = gen_interpretation read_goal_expression | 
| 29211 | 877 | Syntax.parse_prop Attrib.intern_src x; | 
| 28895 | 878 | |
| 38108 | 879 | fun interpret x = gen_interpret cert_goal_expression (K I) (K I) x; | 
| 880 | fun interpret_cmd x = gen_interpret read_goal_expression | |
| 881 | Syntax.parse_prop Attrib.intern_src x; | |
| 882 | ||
| 28895 | 883 | end; | 
| 28903 
b3fc3a62247a
Intro_locales_tac to simplify goals involving locale predicates.
 ballarin parents: 
28902diff
changeset | 884 | |
| 29018 | 885 | |
| 32074 | 886 | (** Interpretation between locales: declaring sublocale relationships **) | 
| 887 | ||
| 888 | local | |
| 889 | ||
| 41270 | 890 | fun note_eqns_dependency target deps witss attrss eqns export export' ctxt = | 
| 891 | let | |
| 892 | val facts = | |
| 893 | (attrss ~~ map (fn eqn => [([Morphism.thm (export' $> export) eqn], [])]) eqns); | |
| 45588 | 894 | val eqns' = ctxt | 
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47068diff
changeset | 895 | |> Attrib.local_notes Thm.lemmaK facts | 
| 45588 | 896 | |-> (fn facts => `(fn ctxt => meta_rewrite ctxt (maps snd facts))) | 
| 41270 | 897 | |> fst; (* FIXME duplication to add_thmss *) | 
| 898 | in | |
| 899 | ctxt | |
| 45589 
bb944d58ac19
simplified Locale.add_thmss, after partial evaluation of attributes;
 wenzelm parents: 
45588diff
changeset | 900 | |> Locale.add_thmss target Thm.lemmaK (Attrib.partial_evaluation ctxt facts) | 
| 42360 | 901 | |> Proof_Context.background_theory (fold (fn ((dep, morph), wits) => | 
| 45588 | 902 | fn thy => | 
| 41270 | 903 | Locale.add_dependency target | 
| 46856 | 904 | (dep, morph $> Element.satisfy_morphism (map (Element.transform_witness export') wits)) | 
| 45588 | 905 | (Element.eq_morphism thy eqns' |> Option.map (rpair true)) | 
| 906 | export thy) (deps ~~ witss)) | |
| 41270 | 907 | end; | 
| 908 | ||
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 909 | fun gen_sublocale prep_expr prep_loc parse_prop prep_attr | 
| 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 | 910 | before_exit raw_target expression equations thy = | 
| 32074 | 911 | let | 
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 912 | val target = prep_loc thy raw_target; | 
| 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 | 913 | val target_ctxt = Named_Target.init before_exit target thy; | 
| 41270 | 914 | val ((propss, deps, export), expr_ctxt) = prep_expr expression target_ctxt; | 
| 43543 
eb8b4851b039
While reading equations of an interpretation, already allow syntax provided by the interpretation base.
 ballarin parents: 
42494diff
changeset | 915 | val eqns = read_with_extended_syntax parse_prop deps expr_ctxt equations; | 
| 41270 | 916 | |
| 45587 | 917 | val attrss = map (apsnd (map (prep_attr thy)) o fst) equations; | 
| 41270 | 918 | val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt; | 
| 919 | val export' = Variable.export_morphism goal_ctxt expr_ctxt; | |
| 920 | ||
| 921 | fun after_qed witss eqns = | |
| 922 | note_eqns_dependency target deps witss attrss eqns export export'; | |
| 923 | ||
| 924 | in Element.witness_proof_eqs after_qed propss eqns goal_ctxt end; | |
| 32074 | 925 | in | 
| 926 | ||
| 41270 | 927 | fun sublocale x = gen_sublocale cert_goal_expression (K I) (K I) (K I) x; | 
| 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 | 928 | fun sublocale_cmd x = | 
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46899diff
changeset | 929 | gen_sublocale read_goal_expression Locale.check Syntax.parse_prop Attrib.intern_src x; | 
| 32074 | 930 | |
| 931 | end; | |
| 932 | ||
| 41435 | 933 | |
| 934 | (** Print the instances that would be activated by an interpretation | |
| 935 | of the expression in the current context (clean = false) or in an | |
| 936 | empty context (clean = true). **) | |
| 937 | ||
| 938 | fun print_dependencies ctxt clean expression = | |
| 939 | let | |
| 940 | val ((_, deps, export), expr_ctxt) = read_goal_expression expression ctxt; | |
| 941 | in | |
| 942 | Locale.print_dependencies expr_ctxt clean export deps | |
| 943 | end; | |
| 944 | ||
| 28993 
829e684b02ef
Interpretation in theories including interaction with subclass relation.
 ballarin parents: 
28951diff
changeset | 945 | end; | 
| 
829e684b02ef
Interpretation in theories including interaction with subclass relation.
 ballarin parents: 
28951diff
changeset | 946 |