src/Pure/Isar/locale.ML
author wenzelm
Wed, 02 Aug 2006 22:27:01 +0200
changeset 20307 a44096c94b3c
parent 20243 8b64a1ea9b1b
child 20317 6e070b33e72b
permissions -rw-r--r--
removed obsolete Drule.frees/vars_of etc.; simplified Assumption/ProofContext.export;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12014
035ab884b9e0 beginnings of new locales (not yet functional);
wenzelm
parents: 11896
diff changeset
     1
(*  Title:      Pure/Isar/locale.ML
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
     3
    Author:     Clemens Ballarin, TU Muenchen; Markus Wenzel, LMU/TU Muenchen
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
     4
12058
cc182b43dd55 fixes: optional typ;
wenzelm
parents: 12046
diff changeset
     5
Locales -- Isar proof contexts as meta-level predicates, with local
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
     6
syntax and implicit structures.
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
     7
19942
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
     8
Draws basic ideas from Florian Kammueller's original version of
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
     9
locales, but uses the richer infrastructure of Isar instead of the raw
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
    10
meta-logic.  Furthermore, structured import of contexts (with merge
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
    11
and rename operations) are provided, as well as type-inference of the
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
    12
signature parts, and predicate definitions of the specification text.
14446
0bc2519e9990 *** empty log message ***
ballarin
parents: 14291
diff changeset
    13
17437
9deaf32c83be interpretation uses primitive goal interface
ballarin
parents: 17412
diff changeset
    14
Interpretation enables the reuse of theorems of locales in other
9deaf32c83be interpretation uses primitive goal interface
ballarin
parents: 17412
diff changeset
    15
contexts, namely those defined by theories, structured proofs and
9deaf32c83be interpretation uses primitive goal interface
ballarin
parents: 17412
diff changeset
    16
locales themselves.
9deaf32c83be interpretation uses primitive goal interface
ballarin
parents: 17412
diff changeset
    17
14446
0bc2519e9990 *** empty log message ***
ballarin
parents: 14291
diff changeset
    18
See also:
0bc2519e9990 *** empty log message ***
ballarin
parents: 14291
diff changeset
    19
0bc2519e9990 *** empty log message ***
ballarin
parents: 14291
diff changeset
    20
[1] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar.
0bc2519e9990 *** empty log message ***
ballarin
parents: 14291
diff changeset
    21
    In Stefano Berardi et al., Types for Proofs and Programs: International
15099
6d8619440ea0 Some comments added.
ballarin
parents: 14981
diff changeset
    22
    Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004.
19942
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
    23
[2] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
    24
    Dependencies between Locales. Technical Report TUM-I0607, Technische
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
    25
    Universitaet Muenchen, 2006.
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
    26
[3] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
    27
    Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108,
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
    28
    pages 31-43, 2006.
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
    29
*)
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
    30
16169
b59202511b8a Locales: new element constrains, parameter renaming with syntax,
ballarin
parents: 16144
diff changeset
    31
(* TODO:
b59202511b8a Locales: new element constrains, parameter renaming with syntax,
ballarin
parents: 16144
diff changeset
    32
- beta-eta normalisation of interpretation parameters
b59202511b8a Locales: new element constrains, parameter renaming with syntax,
ballarin
parents: 16144
diff changeset
    33
- dangling type frees in locales
16620
2a7f46324218 Proper treatment of beta-redexes in witness theorems.
ballarin
parents: 16458
diff changeset
    34
- test subsumption of interpretations when merging theories
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
    35
- var vs. fixes in locale to be interpreted (interpretation in locale)
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
    36
  (implicit locale expressions generated by multiple registrations)
16169
b59202511b8a Locales: new element constrains, parameter renaming with syntax,
ballarin
parents: 16144
diff changeset
    37
*)
b59202511b8a Locales: new element constrains, parameter renaming with syntax,
ballarin
parents: 16144
diff changeset
    38
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
    39
signature LOCALE =
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
    40
sig
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
    41
  datatype expr =
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
    42
    Locale of string |
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
    43
    Rename of expr * (string * mixfix option) option list |
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
    44
    Merge of expr list
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
    45
  val empty: expr
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
    46
  datatype 'a element = Elem of 'a | Expr of expr
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
    47
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
    48
  val intern: theory -> xstring -> string
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
    49
  val extern: theory -> string -> xstring
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
    50
  val init: string -> theory -> Proof.context
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
    51
18795
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
    52
  (* The specification of a locale *)
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
    53
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
    54
  val parameters_of: theory -> string ->
18917
77e18862990f refined signature of locale module
haftmann
parents: 18907
diff changeset
    55
    ((string * typ) * mixfix) list
19276
ac90764bb654 New interface function parameters_of_expr.
ballarin
parents: 19061
diff changeset
    56
  val parameters_of_expr: theory -> expr ->
ac90764bb654 New interface function parameters_of_expr.
ballarin
parents: 19061
diff changeset
    57
    ((string * typ) * mixfix) list
18795
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
    58
  val local_asms_of: theory -> string ->
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
    59
    ((string * Attrib.src list) * term list) list
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
    60
  val global_asms_of: theory -> string ->
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
    61
    ((string * Attrib.src list) * term list) list
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
    62
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
    63
  (* Processing of locale statements *)
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
    64
  val read_context_statement: xstring option -> Element.context element list ->
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
    65
    (string * string list) list list -> Proof.context ->
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
    66
    string option * Proof.context * Proof.context * (term * term list) list list
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
    67
  val cert_context_statement: string option -> Element.context_i element list ->
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
    68
    (term * term list) list list -> Proof.context ->
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
    69
    string option * Proof.context * Proof.context * (term * term list) list list
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
    70
  val read_expr: expr -> Element.context list -> Proof.context ->
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
    71
    Element.context_i list * Proof.context
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
    72
  val cert_expr: expr -> Element.context_i list -> Proof.context ->
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
    73
    Element.context_i list * Proof.context
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
    74
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
    75
  (* Diagnostic functions *)
12758
f6aceb9d4b8e print_locale: allow full body specification;
wenzelm
parents: 12730
diff changeset
    76
  val print_locales: theory -> unit
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
    77
  val print_locale: theory -> bool -> expr -> Element.context list -> unit
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
    78
  val print_global_registrations: bool -> string -> theory -> unit
18617
8928e8722301 added init;
wenzelm
parents: 18569
diff changeset
    79
  val print_local_registrations': bool -> string -> Proof.context -> unit
8928e8722301 added init;
wenzelm
parents: 18569
diff changeset
    80
  val print_local_registrations: bool -> string -> Proof.context -> unit
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
    81
18917
77e18862990f refined signature of locale module
haftmann
parents: 18907
diff changeset
    82
  val add_locale: bool -> bstring -> expr -> Element.context list -> theory
19293
a67b9916c58e Tuned signature of Locale.add_locale(_i).
ballarin
parents: 19282
diff changeset
    83
    -> (string * Proof.context (*body context!*)) * theory
18917
77e18862990f refined signature of locale module
haftmann
parents: 18907
diff changeset
    84
  val add_locale_i: bool -> bstring -> expr -> Element.context_i list -> theory
19293
a67b9916c58e Tuned signature of Locale.add_locale(_i).
ballarin
parents: 19282
diff changeset
    85
    -> (string * Proof.context (*body context!*)) * theory
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
    86
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
    87
  (* Storing results *)
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
    88
  val note_thmss: string -> xstring ->
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
    89
    ((string * Attrib.src list) * (thmref * Attrib.src list) list) list ->
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
    90
    theory -> ((string * thm list) list * (string * thm list) list) * (Proof.context * theory)
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
    91
  val note_thmss_i: string -> string ->
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
    92
    ((string * Attrib.src list) * (thm list * Attrib.src list) list) list ->
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
    93
    theory -> ((string * thm list) list * (string * thm list) list) * (Proof.context * theory)
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
    94
  val add_thmss: string -> string ->
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
    95
    ((string * Attrib.src list) * (thm list * Attrib.src list) list) list ->
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
    96
    Proof.context ->
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
    97
    ((string * thm list) list * (string * thm list) list) * Proof.context
20037
d4102c7cf051 Fixed erroneous check-in.
ballarin
parents: 20035
diff changeset
    98
  val add_term_syntax: string -> (Proof.context -> Proof.context) -> Proof.context -> Proof.context
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
    99
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   100
  val theorem: string -> Method.text option -> (thm list list -> theory -> theory) ->
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
   101
    string * Attrib.src list -> Element.context element list -> Element.statement ->
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   102
    theory -> Proof.state
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   103
  val theorem_i: string -> Method.text option -> (thm list list -> theory -> theory) ->
18907
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
   104
    string * Attrib.src list -> Element.context_i element list -> Element.statement_i ->
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   105
    theory -> Proof.state
17856
0551978bfda5 goal statements: accomodate before_qed argument;
wenzelm
parents: 17756
diff changeset
   106
  val theorem_in_locale: string -> Method.text option ->
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   107
    (thm list list -> thm list list -> theory -> theory) ->
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
   108
    xstring -> string * Attrib.src list -> Element.context element list -> Element.statement ->
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   109
    theory -> Proof.state
17856
0551978bfda5 goal statements: accomodate before_qed argument;
wenzelm
parents: 17756
diff changeset
   110
  val theorem_in_locale_i: string -> Method.text option ->
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   111
    (thm list list -> thm list list -> theory -> theory) ->
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   112
    string -> string * Attrib.src list -> Element.context_i element list ->
18907
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
   113
    Element.statement_i -> theory -> Proof.state
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   114
  val smart_theorem: string -> xstring option ->
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
   115
    string * Attrib.src list -> Element.context element list -> Element.statement ->
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   116
    theory -> Proof.state
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   117
  val interpretation: string * Attrib.src list -> expr -> string option list ->
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   118
    theory -> Proof.state
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   119
  val interpretation_in_locale: xstring * expr -> theory -> Proof.state
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   120
  val interpret: string * Attrib.src list -> expr -> string option list ->
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   121
    bool -> Proof.state -> Proof.state
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   122
end;
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
   123
12289
ec2dafd0a6a9 clarified locale operations (rename, merge);
wenzelm
parents: 12277
diff changeset
   124
structure Locale: LOCALE =
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   125
struct
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   126
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   127
fun legacy_unvarifyT thm =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   128
  let
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   129
    val cT = Thm.ctyp_of (Thm.theory_of_thm thm);
20307
a44096c94b3c removed obsolete Drule.frees/vars_of etc.;
wenzelm
parents: 20243
diff changeset
   130
    val tvars = rev (Drule.fold_terms Term.add_tvars thm []);
a44096c94b3c removed obsolete Drule.frees/vars_of etc.;
wenzelm
parents: 20243
diff changeset
   131
    val tfrees = map (fn ((x, _), S) => SOME (cT (TFree (x, S)))) tvars;
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   132
  in Drule.instantiate' tfrees [] thm end;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   133
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   134
fun legacy_unvarify raw_thm =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   135
  let
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   136
    val thm = legacy_unvarifyT raw_thm;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   137
    val ct = Thm.cterm_of (Thm.theory_of_thm thm);
20307
a44096c94b3c removed obsolete Drule.frees/vars_of etc.;
wenzelm
parents: 20243
diff changeset
   138
    val vars = rev (Drule.fold_terms Term.add_vars thm []);
a44096c94b3c removed obsolete Drule.frees/vars_of etc.;
wenzelm
parents: 20243
diff changeset
   139
    val frees = map (fn ((x, _), T) => SOME (ct (Free (x, T)))) vars;
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   140
  in Drule.instantiate' [] frees thm end;
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   141
20307
a44096c94b3c removed obsolete Drule.frees/vars_of etc.;
wenzelm
parents: 20243
diff changeset
   142
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
   143
(** locale elements and expressions **)
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   144
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   145
datatype ctxt = datatype Element.ctxt;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   146
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
   147
datatype expr =
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
   148
  Locale of string |
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   149
  Rename of expr * (string * mixfix option) option list |
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
   150
  Merge of expr list;
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   151
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
   152
val empty = Merge [];
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
   153
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   154
datatype 'a element =
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   155
  Elem of 'a | Expr of expr;
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
   156
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   157
fun map_elem f (Elem e) = Elem (f e)
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   158
  | map_elem _ (Expr e) = Expr e;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   159
12070
c72fe1edc9e7 proper treatment of local syntax;
wenzelm
parents: 12063
diff changeset
   160
type locale =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   161
 {axiom: Element.witness list,
19942
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
   162
    (* For locales that define predicates this is [A [A]], where A is the locale
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   163
       specification.  Otherwise []. *)
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
   164
  import: expr,                                                     (*dynamic import*)
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   165
  elems: (Element.context_i * stamp) list,
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   166
    (* Static content, neither Fixes nor Constrains elements *)
19278
4d762b77d319 Internal restructuring: local parameters.
ballarin
parents: 19276
diff changeset
   167
  params: ((string * typ) * mixfix) list,                           (*all params*)
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   168
  lparams: string list,                                             (*local params*)
19662
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
   169
  term_syntax: ((Proof.context -> Proof.context) * stamp) list, (* FIXME depend on morphism *)
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   170
  regs: ((string * string list) * Element.witness list) list,
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   171
    (* Registrations: indentifiers and witnesses of locales interpreted in the locale. *)
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   172
  intros: thm list * thm list}
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   173
    (* Introduction rules: of delta predicate and locale predicate. *)
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   174
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   175
(* CB: an internal (Int) locale element was either imported or included,
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   176
   an external (Ext) element appears directly in the locale text. *)
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   177
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   178
datatype ('a, 'b) int_ext = Int of 'a | Ext of 'b;
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   179
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   180
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   181
16736
1e792b32abef Preparations for interpretation of locales in locales.
ballarin
parents: 16620
diff changeset
   182
(** management of registrations in theories and proof contexts **)
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   183
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   184
structure Registrations :
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   185
  sig
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   186
    type T
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   187
    val empty: T
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   188
    val join: T * T -> T
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   189
    val dest: theory -> T ->
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   190
      (term list * ((string * Attrib.src list) * Element.witness list)) list
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   191
    val lookup: theory -> T * term list ->
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   192
      ((string * Attrib.src list) * Element.witness list) option
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   193
    val insert: theory -> term list * (string * Attrib.src list) -> T ->
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   194
      T * (term list * ((string * Attrib.src list) * Element.witness list)) list
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   195
    val add_witness: term list -> Element.witness -> T -> T
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   196
  end =
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   197
struct
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   198
  (* a registration consists of theorems (actually, witnesses) instantiating locale
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   199
     assumptions and prefix and attributes, indexed by parameter instantiation *)
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   200
  type T = ((string * Attrib.src list) * Element.witness list) Termtab.table;
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   201
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   202
  val empty = Termtab.empty;
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   203
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   204
  (* term list represented as single term, for simultaneous matching *)
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   205
  fun termify ts =
18343
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
   206
    Term.list_comb (Const ("", map fastype_of ts ---> propT), ts);
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   207
  fun untermify t =
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   208
    let fun ut (Const _) ts = ts
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   209
          | ut (s $ t) ts = ut s (t::ts)
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   210
    in ut t [] end;
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   211
16620
2a7f46324218 Proper treatment of beta-redexes in witness theorems.
ballarin
parents: 16458
diff changeset
   212
  (* joining of registrations: prefix and attributes of left theory,
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   213
     thms are equal, no attempt to subsumption testing *)
19025
596fb1eb7856 simplified TableFun.join;
wenzelm
parents: 19018
diff changeset
   214
  fun join (r1, r2) = Termtab.join (fn _ => fn (reg, _) => reg) (r1, r2);
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   215
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   216
  fun dest_transfer thy regs =
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   217
    Termtab.dest regs |> map (apsnd (apsnd (map (Element.transfer_witness thy))));
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   218
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   219
  fun dest thy regs = dest_transfer thy regs |> map (apfst untermify);
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   220
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   221
  (* registrations that subsume t *)
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17142
diff changeset
   222
  fun subsumers thy t regs =
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   223
    filter (fn (t', _) => Pattern.matches thy (t', t)) (dest_transfer thy regs);
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   224
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   225
  (* look up registration, pick one that subsumes the query *)
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   226
  fun lookup thy (regs, ts) =
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   227
    let
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   228
      val t = termify ts;
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   229
      val subs = subsumers thy t regs;
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   230
    in (case subs of
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   231
        [] => NONE
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   232
      | ((t', (attn, thms)) :: _) => let
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   233
            val (tinst, inst) = Pattern.match thy (t', t) (Vartab.empty, Vartab.empty);
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   234
            (* thms contain Frees, not Vars *)
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   235
            val tinst' = tinst |> Vartab.dest
19810
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
   236
                 |> map (fn ((x, 0), (_, T)) => (x, Logic.legacy_unvarifyT T))
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   237
                 |> Symtab.make;
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   238
            val inst' = inst |> Vartab.dest
19810
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
   239
                 |> map (fn ((x, 0), (_, t)) => (x, Logic.legacy_unvarify t))
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   240
                 |> Symtab.make;
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   241
          in
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   242
            SOME (attn, map (Element.inst_witness thy (tinst', inst')) thms)
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   243
          end)
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   244
    end;
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   245
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   246
  (* add registration if not subsumed by ones already present,
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   247
     additionally returns registrations that are strictly subsumed *)
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   248
  fun insert thy (ts, attn) regs =
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   249
    let
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   250
      val t = termify ts;
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   251
      val subs = subsumers thy t regs ;
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   252
    in (case subs of
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   253
        [] => let
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   254
                val sups =
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   255
                  filter (fn (t', _) => Pattern.matches thy (t, t')) (dest_transfer thy regs);
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   256
                val sups' = map (apfst untermify) sups
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
   257
              in (Termtab.update (t, (attn, [])) regs, sups') end
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   258
      | _ => (regs, []))
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   259
    end;
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   260
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   261
  (* add witness theorem to registration,
16169
b59202511b8a Locales: new element constrains, parameter renaming with syntax,
ballarin
parents: 16144
diff changeset
   262
     only if instantiation is exact, otherwise exception Option raised *)
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   263
  fun add_witness ts thm regs =
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   264
    let
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   265
      val t = termify ts;
18343
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
   266
      val (x, thms) = the (Termtab.lookup regs t);
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   267
    in
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
   268
      Termtab.update (t, (x, thm::thms)) regs
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   269
    end;
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   270
end;
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   271
16736
1e792b32abef Preparations for interpretation of locales in locales.
ballarin
parents: 16620
diff changeset
   272
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   273
(** theory data **)
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   274
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   275
structure GlobalLocalesData = TheoryDataFun
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   276
(struct
12014
035ab884b9e0 beginnings of new locales (not yet functional);
wenzelm
parents: 11896
diff changeset
   277
  val name = "Isar/locales";
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   278
  type T = NameSpace.T * locale Symtab.table * Registrations.T Symtab.table;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   279
    (* 1st entry: locale namespace,
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   280
       2nd entry: locales of the theory,
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   281
       3rd entry: registrations, indexed by locale name *)
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   282
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   283
  val empty = (NameSpace.empty, Symtab.empty, Symtab.empty);
12063
4c16bdee47d4 added add_locale(_i) and store_thm;
wenzelm
parents: 12058
diff changeset
   284
  val copy = I;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   285
  val extend = I;
12289
ec2dafd0a6a9 clarified locale operations (rename, merge);
wenzelm
parents: 12277
diff changeset
   286
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   287
  fun join_locs _ ({axiom, import, elems, params, lparams, term_syntax, regs, intros}: locale,
19662
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
   288
      {elems = elems', term_syntax = term_syntax', regs = regs', ...}: locale) =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   289
     {axiom = axiom,
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
   290
      import = import,
17496
26535df536ae slight adaptions to library changes
haftmann
parents: 17485
diff changeset
   291
      elems = gen_merge_lists (eq_snd (op =)) elems elems',
16736
1e792b32abef Preparations for interpretation of locales in locales.
ballarin
parents: 16620
diff changeset
   292
      params = params,
19278
4d762b77d319 Internal restructuring: local parameters.
ballarin
parents: 19276
diff changeset
   293
      lparams = lparams,
19662
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
   294
      term_syntax = Library.merge (eq_snd (op =)) (term_syntax, term_syntax'),
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   295
      regs = merge_alists regs regs',
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   296
      intros = intros};
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   297
  fun merge _ ((space1, locs1, regs1), (space2, locs2, regs2)) =
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   298
    (NameSpace.merge (space1, space2), Symtab.join join_locs (locs1, locs2),
19025
596fb1eb7856 simplified TableFun.join;
wenzelm
parents: 19018
diff changeset
   299
     Symtab.join (K Registrations.join) (regs1, regs2));
12289
ec2dafd0a6a9 clarified locale operations (rename, merge);
wenzelm
parents: 12277
diff changeset
   300
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   301
  fun print _ (space, locs, _) =
16346
baa7b5324fc1 NameSpace.extern_table;
wenzelm
parents: 16325
diff changeset
   302
    Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (space, locs)))
12014
035ab884b9e0 beginnings of new locales (not yet functional);
wenzelm
parents: 11896
diff changeset
   303
    |> Pretty.writeln;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   304
end);
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   305
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18698
diff changeset
   306
val _ = Context.add_setup GlobalLocalesData.init;
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   307
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   308
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   309
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   310
(** context data **)
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   311
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   312
structure LocalLocalesData = ProofDataFun
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   313
(struct
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   314
  val name = "Isar/locales";
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   315
  type T = Registrations.T Symtab.table;
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   316
    (* registrations, indexed by locale name *)
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   317
  fun init _ = Symtab.empty;
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   318
  fun print _ _ = ();
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   319
end);
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   320
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18698
diff changeset
   321
val _ = Context.add_setup LocalLocalesData.init;
12289
ec2dafd0a6a9 clarified locale operations (rename, merge);
wenzelm
parents: 12277
diff changeset
   322
12277
2b28d7dd91f5 improved ordering of evaluated elements;
wenzelm
parents: 12273
diff changeset
   323
2b28d7dd91f5 improved ordering of evaluated elements;
wenzelm
parents: 12273
diff changeset
   324
(* access locales *)
2b28d7dd91f5 improved ordering of evaluated elements;
wenzelm
parents: 12273
diff changeset
   325
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   326
val print_locales = GlobalLocalesData.print;
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   327
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   328
val intern = NameSpace.intern o #1 o GlobalLocalesData.get;
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   329
val extern = NameSpace.extern o #1 o GlobalLocalesData.get;
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   330
16144
e339119f4261 renamed cond_extern to extern;
wenzelm
parents: 16105
diff changeset
   331
fun declare_locale name thy =
e339119f4261 renamed cond_extern to extern;
wenzelm
parents: 16105
diff changeset
   332
  thy |> GlobalLocalesData.map (fn (space, locs, regs) =>
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   333
    (Sign.declare_name thy name space, locs, regs));
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   334
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   335
fun put_locale name loc =
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   336
  GlobalLocalesData.map (fn (space, locs, regs) =>
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
   337
    (space, Symtab.update (name, loc) locs, regs));
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
   338
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
   339
fun get_locale thy name = Symtab.lookup (#2 (GlobalLocalesData.get thy)) name;
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   340
12014
035ab884b9e0 beginnings of new locales (not yet functional);
wenzelm
parents: 11896
diff changeset
   341
fun the_locale thy name =
035ab884b9e0 beginnings of new locales (not yet functional);
wenzelm
parents: 11896
diff changeset
   342
  (case get_locale thy name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   343
    SOME loc => loc
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   344
  | NONE => error ("Unknown locale " ^ quote name));
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   345
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
   346
fun change_locale name f thy =
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
   347
  let
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   348
    val {axiom, import, elems, params, lparams, term_syntax, regs, intros} =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   349
        the_locale thy name;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   350
    val (axiom', import', elems', params', lparams', term_syntax', regs', intros') =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   351
      f (axiom, import, elems, params, lparams, term_syntax, regs, intros);
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
   352
  in
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   353
    put_locale name {axiom = axiom', import = import', elems = elems',
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   354
      params = params', lparams = lparams', term_syntax = term_syntax', regs = regs',
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   355
      intros = intros'} thy
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
   356
  end;
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
   357
12046
a404358fd965 locale elements;
wenzelm
parents: 12014
diff changeset
   358
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   359
(* access registrations *)
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   360
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   361
(* Ids of global registrations are varified,
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   362
   Ids of local registrations are not.
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   363
   Thms of registrations are never varified. *)
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   364
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   365
(* retrieve registration from theory or context *)
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   366
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   367
fun gen_get_registrations get thy_of thy_ctxt name =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
   368
  case Symtab.lookup (get thy_ctxt) name of
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   369
      NONE => []
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   370
    | SOME reg => Registrations.dest (thy_of thy_ctxt) reg;
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   371
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   372
val get_global_registrations =
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   373
     gen_get_registrations (#3 o GlobalLocalesData.get) I;
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   374
val get_local_registrations =
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   375
     gen_get_registrations LocalLocalesData.get ProofContext.theory_of;
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   376
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   377
fun gen_get_registration get thy_of thy_ctxt (name, ps) =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
   378
  case Symtab.lookup (get thy_ctxt) name of
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   379
      NONE => NONE
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   380
    | SOME reg => Registrations.lookup (thy_of thy_ctxt) (reg, ps);
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   381
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   382
val get_global_registration =
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   383
     gen_get_registration (#3 o GlobalLocalesData.get) I;
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   384
val get_local_registration =
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   385
     gen_get_registration LocalLocalesData.get ProofContext.theory_of;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   386
18343
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
   387
val test_global_registration = is_some oo get_global_registration;
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
   388
val test_local_registration = is_some oo get_local_registration;
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   389
fun smart_test_registration ctxt id =
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   390
  let
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   391
    val thy = ProofContext.theory_of ctxt;
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   392
  in
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   393
    test_global_registration thy id orelse test_local_registration ctxt id
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   394
  end;
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   395
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   396
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   397
(* add registration to theory or context, ignored if subsumed *)
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   398
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   399
fun gen_put_registration map_data thy_of (name, ps) attn thy_ctxt =
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   400
  map_data (fn regs =>
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   401
    let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   402
      val thy = thy_of thy_ctxt;
18343
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
   403
      val reg = the_default Registrations.empty (Symtab.lookup regs name);
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   404
      val (reg', sups) = Registrations.insert thy (ps, attn) reg;
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   405
      val _ = if not (null sups) then warning
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   406
                ("Subsumed interpretation(s) of locale " ^
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   407
                 quote (extern thy name) ^
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   408
                 "\nby interpretation(s) with the following prefix(es):\n" ^
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   409
                  commas_quote (map (fn (_, ((s, _), _)) => s) sups))
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   410
              else ();
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
   411
    in Symtab.update (name, reg') regs end) thy_ctxt;
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   412
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   413
val put_global_registration =
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   414
     gen_put_registration (fn f =>
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   415
       GlobalLocalesData.map (fn (space, locs, regs) => (space, locs, f regs))) I;
15837
7a567dcd4cda Subsumption of locale interpretations.
ballarin
parents: 15801
diff changeset
   416
val put_local_registration =
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   417
     gen_put_registration LocalLocalesData.map ProofContext.theory_of;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   418
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
   419
fun put_registration_in_locale name id =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   420
  change_locale name (fn (axiom, import, elems, params, lparams, term_syntax, regs, intros) =>
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   421
    (axiom, import, elems, params, lparams, term_syntax, regs @ [(id, [])], intros));
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   422
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   423
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   424
(* add witness theorem to registration in theory or context,
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   425
   ignored if registration not present *)
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   426
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   427
fun add_witness (name, ps) thm =
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   428
  Symtab.map_entry name (Registrations.add_witness ps thm);
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   429
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   430
fun add_global_witness id thm =
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   431
  GlobalLocalesData.map (fn (space, locs, regs) => (space, locs, add_witness id thm regs));
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   432
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   433
val add_local_witness = LocalLocalesData.map oo add_witness;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   434
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
   435
fun add_witness_in_locale name id thm =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   436
  change_locale name (fn (axiom, import, elems, params, lparams, term_syntax, regs, intros) =>
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   437
    let
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   438
      fun add (id', thms) =
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
   439
        if id = id' then (id', thm :: thms) else (id', thms);
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   440
    in (axiom, import, elems, params, lparams, term_syntax, map add regs, intros) end);
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   441
14215
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
   442
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   443
(* printing of registrations *)
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   444
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
   445
fun gen_print_registrations get_regs mk_ctxt msg show_wits loc thy_ctxt =
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   446
  let
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   447
    val ctxt = mk_ctxt thy_ctxt;
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   448
    val thy = ProofContext.theory_of ctxt;
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   449
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   450
    val prt_term = Pretty.quote o ProofContext.pretty_term ctxt;
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   451
    fun prt_inst ts =
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   452
        Pretty.enclose "(" ")" (Pretty.breaks (map prt_term ts));
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   453
    fun prt_attn (prfx, atts) =
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   454
        Pretty.breaks (Pretty.str prfx :: Args.pretty_attribs ctxt atts);
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   455
    fun prt_witns witns = Pretty.enclose "[" "]"
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   456
      (Pretty.breaks (map (prt_term o Element.witness_prop) witns));
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   457
    fun prt_reg (ts, (("", []), witns)) =
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
   458
        if show_wits
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   459
          then Pretty.block [prt_inst ts, Pretty.fbrk, prt_witns witns]
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   460
          else prt_inst ts
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   461
      | prt_reg (ts, (attn, witns)) =
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
   462
        if show_wits
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   463
          then Pretty.block ((prt_attn attn @
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   464
            [Pretty.str ":", Pretty.brk 1, prt_inst ts, Pretty.fbrk,
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   465
              prt_witns witns]))
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   466
          else Pretty.block ((prt_attn attn @
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   467
            [Pretty.str ":", Pretty.brk 1, prt_inst ts]));
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   468
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   469
    val loc_int = intern thy loc;
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   470
    val regs = get_regs thy_ctxt;
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
   471
    val loc_regs = Symtab.lookup regs loc_int;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   472
  in
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   473
    (case loc_regs of
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   474
        NONE => Pretty.str ("no interpretations in " ^ msg)
15763
b901a127ac73 Interpretation supports statically scoped attributes; documentation.
ballarin
parents: 15749
diff changeset
   475
      | SOME r => let
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
   476
            val r' = Registrations.dest thy r;
15763
b901a127ac73 Interpretation supports statically scoped attributes; documentation.
ballarin
parents: 15749
diff changeset
   477
            val r'' = Library.sort_wrt (fn (_, ((prfx, _), _)) => prfx) r';
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
   478
          in Pretty.big_list ("interpretations in " ^ msg ^ ":")
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   479
            (map prt_reg r'')
15763
b901a127ac73 Interpretation supports statically scoped attributes; documentation.
ballarin
parents: 15749
diff changeset
   480
          end)
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   481
    |> Pretty.writeln
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   482
  end;
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   483
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   484
val print_global_registrations =
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   485
     gen_print_registrations (#3 o GlobalLocalesData.get)
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   486
       ProofContext.init "theory";
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   487
val print_local_registrations' =
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   488
     gen_print_registrations LocalLocalesData.get
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   489
       I "context";
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
   490
fun print_local_registrations show_wits loc ctxt =
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
   491
  (print_global_registrations show_wits loc (ProofContext.theory_of ctxt);
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
   492
   print_local_registrations' show_wits loc ctxt);
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
   493
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   494
12277
2b28d7dd91f5 improved ordering of evaluated elements;
wenzelm
parents: 12273
diff changeset
   495
(* diagnostics *)
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
   496
12277
2b28d7dd91f5 improved ordering of evaluated elements;
wenzelm
parents: 12273
diff changeset
   497
fun err_in_locale ctxt msg ids =
2b28d7dd91f5 improved ordering of evaluated elements;
wenzelm
parents: 12273
diff changeset
   498
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   499
    val thy = ProofContext.theory_of ctxt;
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   500
    fun prt_id (name, parms) =
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   501
      [Pretty.block (Pretty.breaks (map Pretty.str (extern thy name :: parms)))];
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   502
    val prt_ids = flat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids));
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   503
    val err_msg =
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   504
      if forall (equal "" o #1) ids then msg
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   505
      else msg ^ "\n" ^ Pretty.string_of (Pretty.block
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   506
        (Pretty.str "The error(s) above occurred in locale:" :: Pretty.brk 1 :: prt_ids));
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18671
diff changeset
   507
  in error err_msg end;
12063
4c16bdee47d4 added add_locale(_i) and store_thm;
wenzelm
parents: 12058
diff changeset
   508
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   509
fun err_in_locale' ctxt msg ids' = err_in_locale ctxt msg (map fst ids');
12277
2b28d7dd91f5 improved ordering of evaluated elements;
wenzelm
parents: 12273
diff changeset
   510
2b28d7dd91f5 improved ordering of evaluated elements;
wenzelm
parents: 12273
diff changeset
   511
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   512
fun pretty_ren NONE = Pretty.str "_"
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   513
  | pretty_ren (SOME (x, NONE)) = Pretty.str x
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   514
  | pretty_ren (SOME (x, SOME syn)) =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   515
      Pretty.block [Pretty.str x, Pretty.brk 1, Syntax.pretty_mixfix syn];
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   516
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   517
fun pretty_expr thy (Locale name) = Pretty.str (extern thy name)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   518
  | pretty_expr thy (Rename (expr, xs)) =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   519
      Pretty.block [pretty_expr thy expr, Pretty.brk 1, Pretty.block (map pretty_ren xs |> Pretty.breaks)]
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   520
  | pretty_expr thy (Merge es) =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   521
      Pretty.separate "+" (map (pretty_expr thy) es) |> Pretty.block;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   522
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   523
fun err_in_expr _ msg (Merge []) = error msg
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   524
  | err_in_expr ctxt msg expr =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   525
    error (msg ^ "\n" ^ Pretty.string_of (Pretty.block
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   526
      [Pretty.str "The error(s) above occured in locale expression:", Pretty.brk 1,
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   527
       pretty_expr (ProofContext.theory_of ctxt) expr]));
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   528
12046
a404358fd965 locale elements;
wenzelm
parents: 12014
diff changeset
   529
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   530
(** structured contexts: rename + merge + implicit type instantiation **)
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   531
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   532
(* parameter types *)
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   533
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   534
fun frozen_tvars ctxt Ts =
19914
fde2b5c0b42b Variable.importT_inst;
wenzelm
parents: 19900
diff changeset
   535
  #1 (Variable.importT_inst (map Logic.mk_type Ts) ctxt)
19900
21a99d88d925 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19873
diff changeset
   536
  |> map (fn ((xi, S), T) => (xi, (S, T)));
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   537
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   538
fun unify_frozen ctxt maxidx Ts Us =
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   539
  let
18343
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
   540
    fun paramify NONE i = (NONE, i)
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
   541
      | paramify (SOME T) i = apfst SOME (TypeInfer.paramify_dummies T i);
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   542
18343
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
   543
    val (Ts', maxidx') = fold_map paramify Ts maxidx;
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
   544
    val (Us', maxidx'') = fold_map paramify Us maxidx';
16947
c6a90f04924e Sign.typ_unify;
wenzelm
parents: 16861
diff changeset
   545
    val thy = ProofContext.theory_of ctxt;
14215
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
   546
18190
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
   547
    fun unify (SOME T, SOME U) env = (Sign.typ_unify thy (U, T) env
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
   548
          handle Type.TUNIFY => raise TYPE ("unify_frozen: failed to unify types", [U, T], []))
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
   549
      | unify _ env = env;
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
   550
    val (unifier, _) = fold unify (Ts' ~~ Us') (Vartab.empty, maxidx'');
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   551
    val Vs = map (Option.map (Envir.norm_type unifier)) Us';
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   552
    val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map_filter I Vs));
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   553
  in map (Option.map (Envir.norm_type unifier')) Vs end;
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   554
19932
63bd0eeb4e0d fixed sml/nj value restriction problem
haftmann
parents: 19931
diff changeset
   555
fun params_of elemss = distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst) elemss);
63bd0eeb4e0d fixed sml/nj value restriction problem
haftmann
parents: 19931
diff changeset
   556
fun params_of' elemss = distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst o fst) elemss);
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   557
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
   558
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
   559
(* CB: param_types has the following type:
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   560
  ('a * 'b option) list -> ('a * 'b) list *)
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   561
fun param_types ps = map_filter (fn (_, NONE) => NONE | (x, SOME T) => SOME (x, T)) ps;
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   562
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   563
19932
63bd0eeb4e0d fixed sml/nj value restriction problem
haftmann
parents: 19931
diff changeset
   564
fun merge_syntax ctxt ids ss = Symtab.merge (op = : mixfix * mixfix -> bool) ss
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   565
  handle Symtab.DUPS xs => err_in_locale ctxt
16105
a44801c499cb SML/NJ compatibility.
ballarin
parents: 16103
diff changeset
   566
    ("Conflicting syntax for parameter(s): " ^ commas_quote xs) (map fst ids);
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   567
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   568
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   569
(* Distinction of assumed vs. derived identifiers.
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   570
   The former may have axioms relating assumptions of the context to
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   571
   assumptions of the specification fragment (for locales with
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   572
   predicates).  The latter have witnesses relating assumptions of the
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   573
   specification fragment to assumptions of other (assumed) specification
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   574
   fragments. *)
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   575
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   576
datatype 'a mode = Assumed of 'a | Derived of 'a;
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   577
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   578
fun map_mode f (Assumed x) = Assumed (f x)
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   579
  | map_mode f (Derived x) = Derived (f x);
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   580
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   581
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   582
(* flatten expressions *)
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
   583
12510
172d18ec3b54 proper treatment of internal parameters;
wenzelm
parents: 12502
diff changeset
   584
local
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   585
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   586
fun unify_parms ctxt fixed_parms raw_parmss =
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   587
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   588
    val thy = ProofContext.theory_of ctxt;
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   589
    val maxidx = length raw_parmss;
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   590
    val idx_parmss = (0 upto maxidx - 1) ~~ raw_parmss;
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   591
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   592
    fun varify i = Term.map_type_tfree (fn (a, S) => TVar ((a, i), S));
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   593
    fun varify_parms (i, ps) = map (apsnd (varify i)) (param_types ps);
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   594
    val parms = fixed_parms @ maps varify_parms idx_parmss;
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   595
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   596
    fun unify T U envir = Sign.typ_unify thy (U, T) envir
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   597
      handle Type.TUNIFY =>
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   598
        let val prt = Sign.string_of_typ thy in
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   599
          raise TYPE ("unify_parms: failed to unify types " ^
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   600
            prt U ^ " and " ^ prt T, [U, T], [])
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   601
        end;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   602
    fun unify_list (T :: Us) = fold (unify T) Us
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   603
      | unify_list [] = I;
18952
0388d0b0f3d5 TableFun: renamed xxx_multi to xxx_list;
wenzelm
parents: 18917
diff changeset
   604
    val (unifier, _) = fold unify_list (map #2 (Symtab.dest (Symtab.make_list parms)))
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   605
      (Vartab.empty, maxidx);
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   606
19061
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
   607
    val parms' = map (apsnd (Envir.norm_type unifier)) (distinct (eq_fst (op =)) parms);
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   608
    val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map #2 parms'));
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   609
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   610
    fun inst_parms (i, ps) =
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   611
      foldr Term.add_typ_tfrees [] (map_filter snd ps)
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   612
      |> map_filter (fn (a, S) =>
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   613
          let val T = Envir.norm_type unifier' (TVar ((a, i), S))
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   614
          in if T = TFree (a, S) then NONE else SOME (a, T) end)
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   615
      |> Symtab.make;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   616
  in map inst_parms idx_parmss end;
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   617
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   618
in
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   619
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   620
fun unify_elemss _ _ [] = []
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   621
  | unify_elemss _ [] [elems] = [elems]
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   622
  | unify_elemss ctxt fixed_parms elemss =
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   623
      let
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   624
        val thy = ProofContext.theory_of ctxt;
17756
d4a35f82fbb4 minor tweaks for Poplog/ML;
wenzelm
parents: 17496
diff changeset
   625
        val envs = unify_parms ctxt fixed_parms (map (snd o fst o fst) elemss);
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   626
        fun inst ((((name, ps), mode), elems), env) =
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   627
          (((name, map (apsnd (Option.map (Element.instT_type env))) ps),
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   628
              map_mode (map (Element.instT_witness thy env)) mode),
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   629
            map (Element.instT_ctxt thy env) elems);
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
   630
      in map inst (elemss ~~ envs) end;
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   631
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   632
19810
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
   633
fun renaming xs parms = zip_options parms xs
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
   634
  handle Library.UnequalLengths =>
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
   635
    error ("Too many arguments in renaming: " ^
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
   636
      commas (map (fn NONE => "_" | SOME x => quote (fst x)) xs));
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
   637
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
   638
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   639
(* params_of_expr:
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   640
   Compute parameters (with types and syntax) of locale expression.
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   641
*)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   642
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   643
fun params_of_expr ctxt fixed_params expr (prev_parms, prev_types, prev_syn) =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   644
  let
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   645
    val thy = ProofContext.theory_of ctxt;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   646
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   647
    fun merge_tenvs fixed tenv1 tenv2 =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   648
        let
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   649
          val [env1, env2] = unify_parms ctxt fixed
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   650
                [tenv1 |> Symtab.dest |> map (apsnd SOME),
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   651
                 tenv2 |> Symtab.dest |> map (apsnd SOME)]
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   652
        in
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   653
          Symtab.merge (op =) (Symtab.map (Element.instT_type env1) tenv1,
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   654
            Symtab.map (Element.instT_type env2) tenv2)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   655
        end;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   656
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   657
    fun merge_syn expr syn1 syn2 =
19932
63bd0eeb4e0d fixed sml/nj value restriction problem
haftmann
parents: 19931
diff changeset
   658
        Symtab.merge (op = : mixfix * mixfix -> bool) (syn1, syn2)
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   659
        handle Symtab.DUPS xs => err_in_expr ctxt
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   660
          ("Conflicting syntax for parameter(s): " ^ commas_quote xs) expr;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   661
            
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   662
    fun params_of (expr as Locale name) =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   663
          let
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   664
            val {import, params, ...} = the_locale thy name;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   665
            val parms = map (fst o fst) params;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   666
            val (parms', types', syn') = params_of import;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   667
            val all_parms = merge_lists parms' parms;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   668
            val all_types = merge_tenvs [] types' (params |> map fst |> Symtab.make);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   669
            val all_syn = merge_syn expr syn' (params |> map (apfst fst) |> Symtab.make);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   670
          in (all_parms, all_types, all_syn) end
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   671
      | params_of (expr as Rename (e, xs)) =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   672
          let
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   673
            val (parms', types', syn') = params_of e;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   674
            val ren = renaming xs parms';
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   675
            (* renaming may reduce number of parameters *)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   676
            val new_parms = map (Element.rename ren) parms' |> distinct (op =);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   677
            val ren_syn = syn' |> Symtab.dest |> map (Element.rename_var ren);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   678
            val new_syn = fold (Symtab.insert (op =)) ren_syn Symtab.empty
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   679
                handle Symtab.DUP x =>
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   680
                  err_in_expr ctxt ("Conflicting syntax for parameter: " ^ quote x) expr;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   681
            val syn_types = map (apsnd (fn mx => SOME (Type.freeze_type (#1 (TypeInfer.paramify_dummies (TypeInfer.mixfixT mx) 0))))) (Symtab.dest new_syn);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   682
            val ren_types = types' |> Symtab.dest |> map (apfst (Element.rename ren));
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   683
            val (env :: _) = unify_parms ctxt [] 
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   684
                ((ren_types |> map (apsnd SOME)) :: map single syn_types);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   685
            val new_types = fold (Symtab.insert (op =))
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   686
                (map (apsnd (Element.instT_type env)) ren_types) Symtab.empty;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   687
          in (new_parms, new_types, new_syn) end
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   688
      | params_of (Merge es) =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   689
          fold (fn e => fn (parms, types, syn) =>
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   690
                   let
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   691
                     val (parms', types', syn') = params_of e
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   692
                   in
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   693
                     (merge_lists parms parms', merge_tenvs [] types types',
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   694
                      merge_syn e syn syn')
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   695
                   end)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   696
            es ([], Symtab.empty, Symtab.empty)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   697
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   698
      val (parms, types, syn) = params_of expr;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   699
    in
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   700
      (merge_lists prev_parms parms, merge_tenvs fixed_params prev_types types,
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   701
       merge_syn expr prev_syn syn)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   702
    end;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   703
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   704
fun make_params_ids params = [(("", params), ([], Assumed []))];
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   705
fun make_raw_params_elemss (params, tenv, syn) =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   706
    [((("", map (fn p => (p, Symtab.lookup tenv p)) params), Assumed []),
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   707
      Int [Fixes (map (fn p =>
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   708
        (p, Symtab.lookup tenv p, Symtab.lookup syn p |> the)) params)])];
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   709
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   710
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   711
(* flatten_expr:
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   712
   Extend list of identifiers by those new in locale expression expr.
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   713
   Compute corresponding list of lists of locale elements (one entry per
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   714
   identifier).
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   715
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   716
   Identifiers represent locale fragments and are in an extended form:
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   717
     ((name, ps), (ax_ps, axs))
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   718
   (name, ps) is the locale name with all its parameters.
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   719
   (ax_ps, axs) is the locale axioms with its parameters;
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   720
     axs are always taken from the top level of the locale hierarchy,
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   721
     hence axioms may contain additional parameters from later fragments:
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   722
     ps subset of ax_ps.  axs is either singleton or empty.
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   723
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   724
   Elements are enriched by identifier-like information:
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   725
     (((name, ax_ps), axs), elems)
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   726
   The parameters in ax_ps are the axiom parameters, but enriched by type
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   727
   info: now each entry is a pair of string and typ option.  Axioms are
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   728
   type-instantiated.
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   729
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   730
*)
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   731
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   732
fun flatten_expr ctxt ((prev_idents, prev_syntax), expr) =
12014
035ab884b9e0 beginnings of new locales (not yet functional);
wenzelm
parents: 11896
diff changeset
   733
  let
035ab884b9e0 beginnings of new locales (not yet functional);
wenzelm
parents: 11896
diff changeset
   734
    val thy = ProofContext.theory_of ctxt;
12263
6f2acf10e2a2 locale expressions;
wenzelm
parents: 12143
diff changeset
   735
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   736
    fun rename_parms top ren ((name, ps), (parms, mode)) =
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   737
        ((name, map (Element.rename ren) ps),
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   738
         if top
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   739
         then (map (Element.rename ren) parms,
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   740
               map_mode (map (Element.rename_witness ren)) mode)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   741
         else (parms, mode));
12263
6f2acf10e2a2 locale expressions;
wenzelm
parents: 12143
diff changeset
   742
20167
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   743
    (* add (name, ps) and its registrations, recursively; adjust hyps of witnesses *)
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   744
20167
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   745
    fun add_with_regs ((name, pTs), mode) (wits, ids, visited) =
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   746
        if member (fn (a, (b, _)) => a = b) visited (name, map #1 pTs)
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   747
        then (wits, ids, visited)
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   748
        else
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   749
	  let
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   750
	    val {params, regs, ...} = the_locale thy name;
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   751
	    val pTs' = map #1 params;
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   752
	    val ren = map #1 pTs' ~~ map (fn (x, _) => (x, NONE)) pTs;
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   753
	      (* dummy syntax, since required by rename *)
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   754
	    val pTs'' = map (fn ((p, _), (_, T)) => (p, T)) (pTs ~~ pTs');
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   755
	    val [env] = unify_parms ctxt pTs [map (apsnd SOME) pTs''];
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   756
	      (* propagate parameter types, to keep them consistent *)
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   757
	    val regs' = map (fn ((name, ps), wits) =>
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   758
		((name, map (Element.rename ren) ps),
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   759
		 map (Element.transfer_witness thy) wits)) regs;
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   760
	    val new_regs = regs';
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   761
	    val new_ids = map fst new_regs;
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   762
	    val new_idTs = map (apsnd (map (fn p => (p, (the o AList.lookup (op =) pTs) p)))) new_ids;
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   763
20167
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   764
	    val new_wits = new_regs |> map (#2 #> map
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   765
	      (Element.instT_witness thy env #> Element.rename_witness ren #>
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   766
		Element.satisfy_witness wits));
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   767
	    val new_ids' = map (fn (id, wits) =>
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   768
		(id, ([], Derived wits))) (new_ids ~~ new_wits);
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   769
	    val new_idTs' = map (fn ((n, pTs), (_, ([], mode))) =>
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   770
		((n, pTs), mode)) (new_idTs ~~ new_ids');
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   771
	    val new_id = ((name, map #1 pTs), ([], mode));
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   772
	    val (wits', ids', visited') = fold add_with_regs new_idTs'
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   773
              (wits @ flat new_wits, ids, visited @ [new_id]);
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   774
	  in
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   775
	    (wits', ids' @ [new_id], visited')
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   776
	  end;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   777
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   778
    (* distribute top-level axioms over assumed ids *)
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   779
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   780
    fun axiomify all_ps ((name, parms), (_, Assumed _)) axioms =
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   781
        let
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   782
          val {elems, ...} = the_locale thy name;
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   783
          val ts = maps
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   784
            (fn (Assumes asms, _) => maps (map #1 o #2) asms
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   785
              | _ => [])
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   786
            elems;
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
   787
          val (axs1, axs2) = chop (length ts) axioms;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   788
        in (((name, parms), (all_ps, Assumed axs1)), axs2) end
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   789
      | axiomify all_ps (id, (_, Derived ths)) axioms =
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   790
          ((id, (all_ps, Derived ths)), axioms);
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   791
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   792
    (* identifiers of an expression *)
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   793
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   794
    fun identify top (Locale name) =
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   795
    (* CB: ids_ax is a list of tuples of the form ((name, ps), axs),
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   796
       where name is a locale name, ps a list of parameter names and axs
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   797
       a list of axioms relating to the identifier, axs is empty unless
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   798
       identify at top level (top = true);
14215
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
   799
       parms is accumulated list of parameters *)
12289
ec2dafd0a6a9 clarified locale operations (rename, merge);
wenzelm
parents: 12277
diff changeset
   800
          let
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   801
            val {axiom, import, params, ...} = the_locale thy name;
19278
4d762b77d319 Internal restructuring: local parameters.
ballarin
parents: 19276
diff changeset
   802
            val ps = map (#1 o #1) params;
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   803
            val (ids', parms') = identify false import;
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   804
                (* acyclic import dependencies *)
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   805
20167
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   806
            val (_, ids'', _) = add_with_regs ((name, map #1 params), Assumed []) ([], ids', ids');
fe5fd4fd4114 Strict dfs traversal of imported and registered identifiers.
ballarin
parents: 20092
diff changeset
   807
            val ids_ax = if top then fst (fold_map (axiomify ps) ids'' axiom) else ids'';
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   808
            in (ids_ax, merge_lists parms' ps) end
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   809
      | identify top (Rename (e, xs)) =
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
   810
          let
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   811
            val (ids', parms') = identify top e;
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
   812
            val ren = renaming xs parms'
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18671
diff changeset
   813
              handle ERROR msg => err_in_locale' ctxt msg ids';
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
   814
19061
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
   815
            val ids'' = distinct (eq_fst (op =)) (map (rename_parms top ren) ids');
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   816
            val parms'' = distinct (op =) (maps (#2 o #1) ids'');
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   817
          in (ids'', parms'') end
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   818
      | identify top (Merge es) =
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   819
          fold (fn e => fn (ids, parms) =>
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   820
                   let
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   821
                     val (ids', parms') = identify top e
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   822
                   in
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   823
                     (merge_alists ids ids', merge_lists parms parms')
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   824
                   end)
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   825
            es ([], []);
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   826
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   827
    fun inst_wit all_params (t, th) = let
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   828
         val {hyps, prop, ...} = Thm.rep_thm th;
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16850
diff changeset
   829
         val ps = map (apsnd SOME) (fold Term.add_frees (prop :: hyps) []);
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   830
         val [env] = unify_parms ctxt all_params [ps];
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   831
         val t' = Element.instT_term env t;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   832
         val th' = Element.instT_thm thy env th;
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   833
       in (t', th') end;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   834
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   835
    fun eval all_params tenv syn ((name, params), (locale_params, mode)) =
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   836
      let
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   837
        val {params = ps_mx, elems = elems_stamped, ...} = the_locale thy name;
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   838
        val elems = map fst elems_stamped;
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   839
        val ps = map fst ps_mx;
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   840
        fun lookup_syn x = (case Symtab.lookup syn x of SOME Structure => NONE | opt => opt);
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   841
        val locale_params' = map (fn p => (p, Symtab.lookup tenv p |> the)) locale_params;
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   842
        val mode' = map_mode (map (Element.map_witness (inst_wit all_params))) mode;
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   843
        val ren = map fst ps ~~ map (fn p => (p, lookup_syn p)) params;
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   844
        val [env] = unify_parms ctxt all_params [map (apfst (Element.rename ren) o apsnd SOME) ps];
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   845
        val elems' = elems
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   846
              |> map (Element.rename_ctxt ren)
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   847
              |> map (Element.map_ctxt {var = I, typ = I, term = I, fact = I, attrib = I,
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   848
                   name = NameSpace.qualified (space_implode "_" params)})
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   849
              |> map (Element.instT_ctxt thy env)
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   850
      in (((name, map (apsnd SOME) locale_params'), mode'), elems') end;
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   851
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   852
    (* parameters, their types and syntax *)
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   853
    val (all_params', tenv, syn) = params_of_expr ctxt [] expr ([], Symtab.empty, Symtab.empty);
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   854
    val all_params = map (fn p => (p, Symtab.lookup tenv p |> the)) all_params';
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   855
    (* compute identifiers and syntax, merge with previous ones *)
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   856
    val (ids, _) = identify true expr;
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   857
    val idents = gen_rems (eq_fst (op =)) (ids, prev_idents);
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   858
    val syntax = merge_syntax ctxt ids (syn, prev_syntax);
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   859
    (* type-instantiate elements *)
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
   860
    val final_elemss = map (eval all_params tenv syntax) idents;
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   861
  in ((prev_idents @ idents, syntax), final_elemss) end;
12046
a404358fd965 locale elements;
wenzelm
parents: 12014
diff changeset
   862
12510
172d18ec3b54 proper treatment of internal parameters;
wenzelm
parents: 12502
diff changeset
   863
end;
172d18ec3b54 proper treatment of internal parameters;
wenzelm
parents: 12502
diff changeset
   864
12070
c72fe1edc9e7 proper treatment of local syntax;
wenzelm
parents: 12063
diff changeset
   865
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   866
(* activate elements *)
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
   867
12510
172d18ec3b54 proper treatment of internal parameters;
wenzelm
parents: 12502
diff changeset
   868
local
172d18ec3b54 proper treatment of internal parameters;
wenzelm
parents: 12502
diff changeset
   869
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
   870
fun axioms_export axs _ hyps =
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   871
  Element.satisfy_thm axs
20307
a44096c94b3c removed obsolete Drule.frees/vars_of etc.;
wenzelm
parents: 20243
diff changeset
   872
  #> Drule.implies_intr_list (Library.drop (length axs, hyps));
12263
6f2acf10e2a2 locale expressions;
wenzelm
parents: 12143
diff changeset
   873
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   874
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   875
(* NB: derived ids contain only facts at this stage *)
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   876
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   877
fun activate_elem _ _ ((ctxt, mode), Fixes fixes) =
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
   878
      ((ctxt |> ProofContext.add_fixes_i fixes |> snd, mode), [])
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   879
  | activate_elem _ _ ((ctxt, mode), Constrains _) =
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   880
      ((ctxt, mode), [])
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   881
  | activate_elem ax_in_ctxt _ ((ctxt, Assumed axs), Assumes asms) =
13399
c136276dc847 support locale ``views'' (for cumulative predicates);
wenzelm
parents: 13394
diff changeset
   882
      let
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   883
        val asms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) asms;
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   884
        val ts = maps (map #1 o #2) asms';
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
   885
        val (ps, qs) = chop (length ts) axs;
17856
0551978bfda5 goal statements: accomodate before_qed argument;
wenzelm
parents: 17756
diff changeset
   886
        val (_, ctxt') =
20243
8b64a1ea9b1b renamed ProofContext.fix_frees to Variable.fix_frees;
wenzelm
parents: 20224
diff changeset
   887
          ctxt |> fold Variable.fix_frees ts
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   888
          |> ProofContext.add_assms_i (axioms_export (if ax_in_ctxt then ps else [])) asms';
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   889
      in ((ctxt', Assumed qs), []) end
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   890
  | activate_elem _ _ ((ctxt, Derived ths), Assumes asms) =
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   891
      ((ctxt, Derived ths), [])
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   892
  | activate_elem _ _ ((ctxt, Assumed axs), Defines defs) =
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   893
      let
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   894
        val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs;
19732
1c37d117a25d activate Defines: fix frees;
wenzelm
parents: 19662
diff changeset
   895
        val asms = defs' |> map (fn ((name, atts), (t, ps)) =>
1c37d117a25d activate Defines: fix frees;
wenzelm
parents: 19662
diff changeset
   896
            let val ((c, _), t') = LocalDefs.cert_def ctxt t
1c37d117a25d activate Defines: fix frees;
wenzelm
parents: 19662
diff changeset
   897
            in (t', ((if name = "" then Thm.def_name c else name, atts), [(t', ps)])) end);
17856
0551978bfda5 goal statements: accomodate before_qed argument;
wenzelm
parents: 17756
diff changeset
   898
        val (_, ctxt') =
20243
8b64a1ea9b1b renamed ProofContext.fix_frees to Variable.fix_frees;
wenzelm
parents: 20224
diff changeset
   899
          ctxt |> fold (Variable.fix_frees o #1) asms
19732
1c37d117a25d activate Defines: fix frees;
wenzelm
parents: 19662
diff changeset
   900
          |> ProofContext.add_assms_i LocalDefs.def_export (map #2 asms);
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   901
      in ((ctxt', Assumed axs), []) end
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   902
  | activate_elem _ _ ((ctxt, Derived ths), Defines defs) =
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   903
      ((ctxt, Derived ths), [])
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   904
  | activate_elem _ is_ext ((ctxt, mode), Notes facts) =
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   905
      let
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   906
        val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts;
17856
0551978bfda5 goal statements: accomodate before_qed argument;
wenzelm
parents: 17756
diff changeset
   907
        val (res, ctxt') = ctxt |> ProofContext.note_thmss_i facts';
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   908
      in ((ctxt', mode), if is_ext then res else []) end;
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
   909
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   910
fun activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt =
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
   911
  let
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   912
    val thy = ProofContext.theory_of ctxt;
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
   913
    val ((ctxt', _), res) =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   914
        foldl_map (activate_elem ax_in_ctxt (name = "")) ((ProofContext.qualified_names ctxt, mode), elems)
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18671
diff changeset
   915
      handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)]
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   916
    val ctxt'' = if name = "" then ctxt'
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   917
          else let
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   918
              val ps' = map (fn (n, SOME T) => Free (n, T)) ps;
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   919
              val ctxt'' = put_local_registration (name, ps') ("", []) ctxt'
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   920
            in case mode of
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   921
                Assumed axs =>
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   922
                  fold (add_local_witness (name, ps') o
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
   923
                    Element.assume_witness thy o Element.witness_prop) axs ctxt''
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
   924
              | Derived ths => fold (add_local_witness (name, ps')) ths ctxt''
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   925
            end
16144
e339119f4261 renamed cond_extern to extern;
wenzelm
parents: 16105
diff changeset
   926
  in (ProofContext.restore_naming ctxt ctxt'', res) end;
13399
c136276dc847 support locale ``views'' (for cumulative predicates);
wenzelm
parents: 13394
diff changeset
   927
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   928
fun activate_elemss ax_in_ctxt prep_facts =
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   929
    fold_map (fn (((name, ps), mode), raw_elems) => fn ctxt =>
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   930
      let
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   931
        val elems = map (prep_facts ctxt) raw_elems;
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   932
        val (ctxt', res) = apsnd flat
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   933
            (activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt);
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   934
        val elems' = elems |> map (Element.map_ctxt
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   935
          {name = I, var = I, typ = I, term = I, fact = I, attrib = Args.closure});
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   936
      in (((((name, ps), mode), elems'), res), ctxt') end);
12834
e5bec3268932 added locale_facts(_i) interface (useful for simple ML proof tools);
wenzelm
parents: 12806
diff changeset
   937
12546
wenzelm
parents: 12532
diff changeset
   938
in
wenzelm
parents: 12532
diff changeset
   939
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   940
(* CB: activate_facts prep_facts (ctxt, elemss),
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   941
   where elemss is a list of pairs consisting of identifiers and
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   942
   context elements, extends ctxt by the context elements yielding
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   943
   ctxt' and returns (ctxt', (elemss', facts)).
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   944
   Identifiers in the argument are of the form ((name, ps), axs) and
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   945
   assumptions use the axioms in the identifiers to set up exporters
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   946
   in ctxt'.  elemss' does not contain identifiers and is obtained
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   947
   from elemss and the intermediate context with prep_facts.
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   948
   If read_facts or cert_facts is used for prep_facts, these also remove
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
   949
   the internal/external markers from elemss. *)
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
   950
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   951
fun activate_facts ax_in_ctxt prep_facts (ctxt, args) =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   952
  let val ((elemss, factss), ctxt') = activate_elemss ax_in_ctxt prep_facts args ctxt |>> split_list
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   953
  in (ctxt', (elemss, flat factss)) end;
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   954
12510
172d18ec3b54 proper treatment of internal parameters;
wenzelm
parents: 12502
diff changeset
   955
end;
172d18ec3b54 proper treatment of internal parameters;
wenzelm
parents: 12502
diff changeset
   956
12307
459aa05af6be qualify imported facts;
wenzelm
parents: 12289
diff changeset
   957
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   958
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   959
(** prepare locale elements **)
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   960
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   961
(* expressions *)
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   962
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   963
fun intern_expr thy (Locale xname) = Locale (intern thy xname)
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   964
  | intern_expr thy (Merge exprs) = Merge (map (intern_expr thy) exprs)
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
   965
  | intern_expr thy (Rename (expr, xs)) = Rename (intern_expr thy expr, xs);
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   966
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   967
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   968
(* propositions and bindings *)
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
   969
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   970
(* flatten (ctxt, prep_expr) ((ids, syn), expr)
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
   971
   normalises expr (which is either a locale
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
   972
   expression or a single context element) wrt.
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
   973
   to the list ids of already accumulated identifiers.
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
   974
   It returns ((ids', syn'), elemss) where ids' is an extension of ids
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
   975
   with identifiers generated for expr, and elemss is the list of
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   976
   context elements generated from expr.
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   977
   syn and syn' are symtabs mapping parameter names to their syntax.  syn'
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   978
   is an extension of syn.
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   979
   For details, see flatten_expr.
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   980
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   981
   Additionally, for a locale expression, the elems are grouped into a single
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   982
   Int; individual context elements are marked Ext.  In this case, the
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   983
   identifier-like information of the element is as follows:
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   984
   - for Fixes: (("", ps), []) where the ps have type info NONE
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   985
   - for other elements: (("", []), []).
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   986
   The implementation of activate_facts relies on identifier names being
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
   987
   empty strings for external elements.
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
   988
*)
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
   989
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   990
fun flatten (ctxt, _) ((ids, syn), Elem (Fixes fixes)) = let
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   991
        val ids' = ids @ [(("", map #1 fixes), ([], Assumed []))]
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   992
      in
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   993
        ((ids',
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   994
         merge_syntax ctxt ids'
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   995
           (syn, Symtab.make (map (fn fx => (#1 fx, #3 fx)) fixes))
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   996
           handle Symtab.DUPS xs => err_in_locale ctxt
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
   997
             ("Conflicting syntax for parameters: " ^ commas_quote xs)
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
   998
             (map #1 ids')),
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
   999
         [((("", map (rpair NONE o #1) fixes), Assumed []), Ext (Fixes fixes))])
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
  1000
      end
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
  1001
  | flatten _ ((ids, syn), Elem elem) =
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1002
      ((ids @ [(("", []), ([], Assumed []))], syn), [((("", []), Assumed []), Ext elem)])
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
  1003
  | flatten (ctxt, prep_expr) ((ids, syn), Expr expr) =
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
  1004
      apsnd (map (apsnd Int)) (flatten_expr ctxt ((ids, syn), prep_expr expr));
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1005
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1006
local
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1007
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1008
local
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1009
12727
330cb92aaea3 unify_frozen: proper use of maxidx';
wenzelm
parents: 12711
diff changeset
  1010
fun declare_int_elem (ctxt, Fixes fixes) =
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1011
      (ctxt |> ProofContext.add_fixes_i (map (fn (x, T, mx) =>
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1012
        (x, Option.map (Term.map_type_tfree (TypeInfer.param 0)) T, mx)) fixes) |> snd, [])
12727
330cb92aaea3 unify_frozen: proper use of maxidx';
wenzelm
parents: 12711
diff changeset
  1013
  | declare_int_elem (ctxt, _) = (ctxt, []);
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1014
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1015
fun declare_ext_elem prep_vars (ctxt, Fixes fixes) =
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1016
      let val (vars, _) = prep_vars fixes ctxt
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1017
      in (ctxt |> ProofContext.add_fixes_i vars |> snd, []) end
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1018
  | declare_ext_elem prep_vars (ctxt, Constrains csts) =
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1019
      let val (_, ctxt') = prep_vars (map (fn (x, T) => (x, SOME T, NoSyn)) csts) ctxt
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1020
      in (ctxt', []) end
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1021
  | declare_ext_elem _ (ctxt, Assumes asms) = (ctxt, map #2 asms)
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
  1022
  | declare_ext_elem _ (ctxt, Defines defs) = (ctxt, map (fn (_, (t, ps)) => [(t, ps)]) defs)
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1023
  | declare_ext_elem _ (ctxt, Notes facts) = (ctxt, []);
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1024
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1025
fun declare_elems prep_vars (ctxt, (((name, ps), Assumed _), elems)) =
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1026
    let val (ctxt', propps) =
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1027
      (case elems of
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1028
        Int es => foldl_map declare_int_elem (ctxt, es)
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1029
      | Ext e => foldl_map (declare_ext_elem prep_vars) (ctxt, [e]))
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18671
diff changeset
  1030
      handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)]
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1031
    in (ctxt', propps) end
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1032
  | declare_elems _ (ctxt, ((_, Derived _), elems)) = (ctxt, []);
12727
330cb92aaea3 unify_frozen: proper use of maxidx';
wenzelm
parents: 12711
diff changeset
  1033
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1034
in
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1035
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1036
fun declare_elemss prep_vars fixed_params raw_elemss ctxt =
12727
330cb92aaea3 unify_frozen: proper use of maxidx';
wenzelm
parents: 12711
diff changeset
  1037
  let
14215
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
  1038
    (* CB: fix of type bug of goal in target with context elements.
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
  1039
       Parameters new in context elements must receive types that are
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
  1040
       distinct from types of parameters in target (fixed_params).  *)
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
  1041
    val ctxt_with_fixed =
19900
21a99d88d925 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19873
diff changeset
  1042
      fold Variable.declare_term (map Free fixed_params) ctxt;
12727
330cb92aaea3 unify_frozen: proper use of maxidx';
wenzelm
parents: 12711
diff changeset
  1043
    val int_elemss =
330cb92aaea3 unify_frozen: proper use of maxidx';
wenzelm
parents: 12711
diff changeset
  1044
      raw_elemss
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1045
      |> map_filter (fn (id, Int es) => SOME (id, es) | _ => NONE)
14215
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
  1046
      |> unify_elemss ctxt_with_fixed fixed_params;
12727
330cb92aaea3 unify_frozen: proper use of maxidx';
wenzelm
parents: 12711
diff changeset
  1047
    val (_, raw_elemss') =
330cb92aaea3 unify_frozen: proper use of maxidx';
wenzelm
parents: 12711
diff changeset
  1048
      foldl_map (fn ((_, es) :: elemss, (id, Int _)) => (elemss, (id, Int es)) | x => x)
330cb92aaea3 unify_frozen: proper use of maxidx';
wenzelm
parents: 12711
diff changeset
  1049
        (int_elemss, raw_elemss);
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1050
  in foldl_map (declare_elems prep_vars) (ctxt, raw_elemss') end;
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1051
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1052
end;
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1053
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1054
local
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1055
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1056
val norm_term = Envir.beta_norm oo Term.subst_atomic;
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1057
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1058
fun abstract_thm thy eq =
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1059
  Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def;
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
  1060
18190
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1061
fun bind_def ctxt (name, ps) eq (xs, env, ths) =
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1062
  let
18831
f1315d118059 LocalDefs;
wenzelm
parents: 18806
diff changeset
  1063
    val ((y, T), b) = LocalDefs.abs_def eq;
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1064
    val b' = norm_term env b;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1065
    val th = abstract_thm (ProofContext.theory_of ctxt) eq;
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1066
    fun err msg = err_in_locale ctxt (msg ^ ": " ^ quote y) [(name, map fst ps)];
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1067
  in
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1068
    conditional (exists (equal y o #1) xs) (fn () =>
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1069
      err "Attempt to define previously specified variable");
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1070
    conditional (exists (fn (Free (y', _), _) => y = y' | _ => false) env) (fn () =>
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1071
      err "Attempt to redefine variable");
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16850
diff changeset
  1072
    (Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths)
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1073
  end;
12575
34985eee55b1 fixed inst_thm: proper domain of env;
wenzelm
parents: 12546
diff changeset
  1074
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1075
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1076
(* CB: for finish_elems (Int and Ext),
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1077
   extracts specification, only of assumed elements *)
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1078
18190
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1079
fun eval_text _ _ _ (Fixes _) text = text
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1080
  | eval_text _ _ _ (Constrains _) text = text
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1081
  | eval_text _ (_, Assumed _) is_ext (Assumes asms)
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1082
        (((exts, exts'), (ints, ints')), (xs, env, defs)) =
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1083
      let
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1084
        val ts = maps (map #1 o #2) asms;
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1085
        val ts' = map (norm_term env) ts;
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1086
        val spec' =
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1087
          if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1088
          else ((exts, exts'), (ints @ ts, ints' @ ts'));
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16850
diff changeset
  1089
      in (spec', (fold Term.add_frees ts' xs, env, defs)) end
18190
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1090
  | eval_text _ (_, Derived _) _ (Assumes _) text = text
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1091
  | eval_text ctxt (id, Assumed _) _ (Defines defs) (spec, binds) =
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1092
      (spec, fold (bind_def ctxt id o #1 o #2) defs binds)
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1093
  | eval_text _ (_, Derived _) _ (Defines _) text = text
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1094
  | eval_text _ _ _ (Notes _) text = text;
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1095
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1096
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1097
(* for finish_elems (Int),
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1098
   remove redundant elements of derived identifiers,
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1099
   turn assumptions and definitions into facts,
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1100
   adjust hypotheses of facts using witnesses *)
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1101
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
  1102
fun finish_derived _ _ (Assumed _) (Fixes fixes) = SOME (Fixes fixes)
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
  1103
  | finish_derived _ _ (Assumed _) (Constrains csts) = SOME (Constrains csts)
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
  1104
  | finish_derived _ _ (Assumed _) (Assumes asms) = SOME (Assumes asms)
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
  1105
  | finish_derived _ _ (Assumed _) (Defines defs) = SOME (Defines defs)
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
  1106
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1107
  | finish_derived _ _ (Derived _) (Fixes _) = NONE
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1108
  | finish_derived _ _ (Derived _) (Constrains _) = NONE
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1109
  | finish_derived sign wits (Derived _) (Assumes asms) = asms
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
  1110
      |> map (apsnd (map (fn (a, _) => ([Thm.assume (cterm_of sign a)], []))))
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1111
      |> Notes |> Element.map_ctxt_values I I (Element.satisfy_thm wits) |> SOME
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1112
  | finish_derived sign wits (Derived _) (Defines defs) = defs
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
  1113
      |> map (apsnd (fn (d, _) => [([Thm.assume (cterm_of sign d)], [])]))
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1114
      |> Notes |> Element.map_ctxt_values I I (Element.satisfy_thm wits) |> SOME
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1115
17096
8327b71282ce Improved generation of witnesses in interpretation.
ballarin
parents: 17033
diff changeset
  1116
  | finish_derived _ wits _ (Notes facts) = (Notes facts)
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1117
      |> Element.map_ctxt_values I I (Element.satisfy_thm wits) |> SOME;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1118
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1119
(* CB: for finish_elems (Ext) *)
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1120
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1121
fun closeup _ false elem = elem
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1122
  | closeup ctxt true elem =
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1123
      let
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1124
        fun close_frees t =
19900
21a99d88d925 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19873
diff changeset
  1125
          let val frees = rev (filter_out (Variable.is_fixed ctxt o #1) (Term.add_frees t []))
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1126
          in Term.list_all_free (frees, t) end;
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1127
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1128
        fun no_binds [] = []
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18671
diff changeset
  1129
          | no_binds _ = error "Illegal term bindings in locale element";
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1130
      in
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1131
        (case elem of
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1132
          Assumes asms => Assumes (asms |> map (fn (a, propps) =>
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
  1133
            (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1134
        | Defines defs => Defines (defs |> map (fn (a, (t, ps)) =>
18831
f1315d118059 LocalDefs;
wenzelm
parents: 18806
diff changeset
  1135
            (a, (close_frees (#2 (LocalDefs.cert_def ctxt t)), no_binds ps))))
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1136
        | e => e)
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1137
      end;
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1138
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
  1139
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1140
fun finish_ext_elem parms _ (Fixes fixes, _) = Fixes (map (fn (x, _, mx) =>
17271
2756a73f63a5 introduced some new-style AList operations
haftmann
parents: 17257
diff changeset
  1141
      (x, AList.lookup (op =) parms x, mx)) fixes)
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1142
  | finish_ext_elem parms _ (Constrains _, _) = Constrains []
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1143
  | finish_ext_elem _ close (Assumes asms, propp) =
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1144
      close (Assumes (map #1 asms ~~ propp))
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1145
  | finish_ext_elem _ close (Defines defs, propp) =
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
  1146
      close (Defines (map #1 defs ~~ map (fn [(t, ps)] => (t, ps)) propp))
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1147
  | finish_ext_elem _ _ (Notes facts, _) = Notes facts;
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1148
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1149
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1150
(* CB: finish_parms introduces type info from parms to identifiers *)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
  1151
(* CB: only needed for types that have been NONE so far???
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1152
   If so, which are these??? *)
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1153
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1154
fun finish_parms parms (((name, ps), mode), elems) =
19932
63bd0eeb4e0d fixed sml/nj value restriction problem
haftmann
parents: 19931
diff changeset
  1155
  (((name, map (fn (x, _) => (x, AList.lookup (op = : string * string -> bool) parms x)) ps), mode), elems);
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1156
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1157
fun finish_elems ctxt parms _ ((text, wits), ((id, Int e), _)) =
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1158
      let
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1159
        val [(id' as (_, mode), es)] = unify_elemss ctxt parms [(id, e)];
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1160
        val wits' = case mode of Assumed _ => wits | Derived ths => wits @ ths;
18190
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1161
        val text' = fold (eval_text ctxt id' false) es text;
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1162
        val es' = map_filter
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1163
              (finish_derived (ProofContext.theory_of ctxt) wits' mode) es;
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1164
      in ((text', wits'), (id', map Int es')) end
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1165
  | finish_elems ctxt parms do_close ((text, wits), ((id, Ext e), [propp])) =
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1166
      let
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1167
        val e' = finish_ext_elem parms (closeup ctxt do_close) (e, propp);
18190
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1168
        val text' = eval_text ctxt id true e' text;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1169
      in ((text', wits), (id, [Ext e'])) end
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1170
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1171
in
12510
172d18ec3b54 proper treatment of internal parameters;
wenzelm
parents: 12502
diff changeset
  1172
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1173
(* CB: only called by prep_elemss *)
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1174
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1175
fun finish_elemss ctxt parms do_close =
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1176
  foldl_map (apsnd (finish_parms parms) o finish_elems ctxt parms do_close);
12839
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1177
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1178
end;
584a3e0b00f2 reorganized code for predicate text;
wenzelm
parents: 12834
diff changeset
  1179
16736
1e792b32abef Preparations for interpretation of locales in locales.
ballarin
parents: 16620
diff changeset
  1180
19942
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1181
(* Remove duplicate Defines elements: temporary workaround to fix Afp/Category. *)
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1182
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1183
fun defs_ord (defs1, defs2) =
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1184
    list_ord (fn ((_, (d1, _)), (_, (d2, _))) =>
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1185
      Term.fast_term_ord (d1, d2)) (defs1, defs2);
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1186
structure Defstab =
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1187
    TableFun(type key = ((string * Attrib.src list) * (term * term list)) list
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1188
        val ord = defs_ord);
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1189
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1190
fun rem_dup_defs es ds =
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1191
    fold_map (fn e as (Defines defs) => (fn ds =>
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1192
                 if Defstab.defined ds defs
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1193
                 then (Defines [], ds)
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1194
                 else (e, Defstab.update (defs, ()) ds))
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1195
               | e => (fn ds => (e, ds))) es ds;
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1196
fun rem_dup_elemss (Int es) ds = apfst Int (rem_dup_defs es ds)
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1197
  | rem_dup_elemss (Ext e) ds = (Ext e, ds);
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1198
fun rem_dup_defines raw_elemss =
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1199
    fold_map (fn (id as (_, (Assumed _)), es) => (fn ds =>
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1200
                     apfst (pair id) (rem_dup_elemss es ds))
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1201
               | (id as (_, (Derived _)), es) => (fn ds =>
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1202
                     ((id, es), ds))) raw_elemss Defstab.empty |> #1;
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1203
16736
1e792b32abef Preparations for interpretation of locales in locales.
ballarin
parents: 16620
diff changeset
  1204
(* CB: type inference and consistency checks for locales.
1e792b32abef Preparations for interpretation of locales in locales.
ballarin
parents: 16620
diff changeset
  1205
1e792b32abef Preparations for interpretation of locales in locales.
ballarin
parents: 16620
diff changeset
  1206
   Works by building a context (through declare_elemss), extracting the
1e792b32abef Preparations for interpretation of locales in locales.
ballarin
parents: 16620
diff changeset
  1207
   required information and adjusting the context elements (finish_elemss).
1e792b32abef Preparations for interpretation of locales in locales.
ballarin
parents: 16620
diff changeset
  1208
   Can also universally close free vars in assms and defs.  This is only
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1209
   needed for Ext elements and controlled by parameter do_close.
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1210
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1211
   Only elements of assumed identifiers are considered.
16736
1e792b32abef Preparations for interpretation of locales in locales.
ballarin
parents: 16620
diff changeset
  1212
*)
15127
2550a5578d39 Disallowed "includes" in locale declarations.
ballarin
parents: 15104
diff changeset
  1213
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1214
fun prep_elemss prep_vars prepp do_close context fixed_params raw_elemss raw_concl =
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1215
  let
15127
2550a5578d39 Disallowed "includes" in locale declarations.
ballarin
parents: 15104
diff changeset
  1216
    (* CB: contexts computed in the course of this function are discarded.
2550a5578d39 Disallowed "includes" in locale declarations.
ballarin
parents: 15104
diff changeset
  1217
       They are used for type inference and consistency checks only. *)
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1218
    (* CB: fixed_params are the parameters (with types) of the target locale,
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1219
       empty list if there is no target. *)
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1220
    (* CB: raw_elemss are list of pairs consisting of identifiers and
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1221
       context elements, the latter marked as internal or external. *)
19942
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1222
    val raw_elemss = rem_dup_defines raw_elemss;
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1223
    val (raw_ctxt, raw_proppss) = declare_elemss prep_vars fixed_params raw_elemss context;
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1224
    (* CB: raw_ctxt is context with additional fixed variables derived from
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1225
       the fixes elements in raw_elemss,
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1226
       raw_proppss contains assumptions and definitions from the
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1227
       external elements in raw_elemss. *)
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1228
    fun prep_prop raw_propp (raw_ctxt, raw_concl)  =
18450
e57731ba01dd discontinued unflat in favour of burrow and burrow_split
haftmann
parents: 18421
diff changeset
  1229
      let
e57731ba01dd discontinued unflat in favour of burrow and burrow_split
haftmann
parents: 18421
diff changeset
  1230
        (* CB: add type information from fixed_params to context (declare_term) *)
e57731ba01dd discontinued unflat in favour of burrow and burrow_split
haftmann
parents: 18421
diff changeset
  1231
        (* CB: process patterns (conclusion and external elements only) *)
e57731ba01dd discontinued unflat in favour of burrow and burrow_split
haftmann
parents: 18421
diff changeset
  1232
        val (ctxt, all_propp) =
19900
21a99d88d925 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19873
diff changeset
  1233
          prepp (fold Variable.declare_term (map Free fixed_params) raw_ctxt, raw_concl @ raw_propp);
18450
e57731ba01dd discontinued unflat in favour of burrow and burrow_split
haftmann
parents: 18421
diff changeset
  1234
        (* CB: add type information from conclusion and external elements to context *)
19900
21a99d88d925 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19873
diff changeset
  1235
        val ctxt = fold Variable.declare_term (maps (map fst) all_propp) ctxt;
18450
e57731ba01dd discontinued unflat in favour of burrow and burrow_split
haftmann
parents: 18421
diff changeset
  1236
        (* CB: resolve schematic variables (patterns) in conclusion and external elements. *)
e57731ba01dd discontinued unflat in favour of burrow and burrow_split
haftmann
parents: 18421
diff changeset
  1237
        val all_propp' = map2 (curry (op ~~))
e57731ba01dd discontinued unflat in favour of burrow and burrow_split
haftmann
parents: 18421
diff changeset
  1238
          (#1 (#2 (ProofContext.bind_propp_schematic_i (ctxt, all_propp)))) (map (map snd) all_propp);
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1239
        val (concl, propp) = chop (length raw_concl) all_propp';
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1240
      in (propp, (ctxt, concl)) end
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1241
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1242
    val (proppss, (ctxt, concl)) =
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1243
      (fold_burrow o fold_burrow) prep_prop raw_proppss (raw_ctxt, raw_concl);
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
  1244
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1245
    (* CB: obtain all parameters from identifier part of raw_elemss *)
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1246
    val xs = map #1 (params_of' raw_elemss);
12727
330cb92aaea3 unify_frozen: proper use of maxidx';
wenzelm
parents: 12711
diff changeset
  1247
    val typing = unify_frozen ctxt 0
19900
21a99d88d925 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19873
diff changeset
  1248
      (map (Variable.default_type raw_ctxt) xs)
21a99d88d925 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19873
diff changeset
  1249
      (map (Variable.default_type ctxt) xs);
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1250
    val parms = param_types (xs ~~ typing);
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1251
    (* CB: parms are the parameters from raw_elemss, with correct typing. *)
12273
7fb9840d358d beginnings of actual locale expressions;
wenzelm
parents: 12263
diff changeset
  1252
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1253
    (* CB: extract information from assumes and defines elements
16169
b59202511b8a Locales: new element constrains, parameter renaming with syntax,
ballarin
parents: 16144
diff changeset
  1254
       (fixes, constrains and notes in raw_elemss don't have an effect on
b59202511b8a Locales: new element constrains, parameter renaming with syntax,
ballarin
parents: 16144
diff changeset
  1255
       text and elemss), compute final form of context elements. *)
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1256
    val ((text, _), elemss) = finish_elemss ctxt parms do_close
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1257
      ((((([], []), ([], [])), ([], [], [])), []), raw_elemss ~~ proppss);
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1258
    (* CB: text has the following structure:
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1259
           (((exts, exts'), (ints, ints')), (xs, env, defs))
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1260
       where
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1261
         exts: external assumptions (terms in external assumes elements)
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1262
         exts': dito, normalised wrt. env
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1263
         ints: internal assumptions (terms in internal assumes elements)
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1264
         ints': dito, normalised wrt. env
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1265
         xs: the free variables in exts' and ints' and rhss of definitions,
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1266
           this includes parameters except defined parameters
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1267
         env: list of term pairs encoding substitutions, where the first term
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1268
           is a free variable; substitutions represent defines elements and
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1269
           the rhs is normalised wrt. the previous env
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1270
         defs: theorems representing the substitutions from defines elements
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1271
           (thms are normalised wrt. env).
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1272
       elemss is an updated version of raw_elemss:
16169
b59202511b8a Locales: new element constrains, parameter renaming with syntax,
ballarin
parents: 16144
diff changeset
  1273
         - type info added to Fixes and modified in Constrains
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1274
         - axiom and definition statement replaced by corresponding one
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1275
           from proppss in Assumes and Defines
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1276
         - Facts unchanged
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1277
       *)
13308
1dbed9ea764b clarified text content of locale body;
wenzelm
parents: 13297
diff changeset
  1278
  in ((parms, elemss, concl), text) end;
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
  1279
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
  1280
in
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
  1281
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1282
fun read_elemss x = prep_elemss ProofContext.read_vars ProofContext.read_propp_schematic x;
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1283
fun cert_elemss x = prep_elemss ProofContext.cert_vars ProofContext.cert_propp_schematic x;
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1284
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1285
end;
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1286
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1287
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
  1288
(* facts and attributes *)
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1289
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1290
local
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1291
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18671
diff changeset
  1292
fun prep_name name =
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18671
diff changeset
  1293
  if NameSpace.is_qualified name then error ("Illegal qualified name: " ^ quote name)
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
  1294
  else name;
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1295
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
  1296
fun prep_facts _ _ ctxt (Int elem) =
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1297
      Element.map_ctxt_values I I (Thm.transfer (ProofContext.theory_of ctxt)) elem
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1298
  | prep_facts get intern ctxt (Ext elem) = elem |> Element.map_ctxt
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
  1299
     {var = I, typ = I, term = I,
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18671
diff changeset
  1300
      name = prep_name,
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
  1301
      fact = get ctxt,
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1302
      attrib = Args.assignable o intern (ProofContext.theory_of ctxt)};
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1303
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1304
in
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1305
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
  1306
fun read_facts x = prep_facts ProofContext.get_thms Attrib.intern_src x;
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
  1307
fun cert_facts x = prep_facts (K I) (K I) x;
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1308
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1309
end;
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1310
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1311
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1312
(* Get the specification of a locale *)
18795
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1313
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1314
(*The global specification is made from the parameters and global
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1315
  assumptions, the local specification from the parameters and the
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1316
  local assumptions.*)
18795
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1317
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1318
local
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1319
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1320
fun gen_asms_of get thy name =
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1321
  let
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1322
    val ctxt = ProofContext.init thy;
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1323
    val (_, raw_elemss) = flatten (ctxt, I) (([], Symtab.empty), Expr (Locale name));
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1324
    val ((_, elemss, _), _) = read_elemss false ctxt [] raw_elemss [];
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1325
  in
18890
34a9e466201f *_asms_of fixed.
ballarin
parents: 18831
diff changeset
  1326
    elemss |> get
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1327
      |> maps (fn (_, es) => map (fn Int e => e) es)
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1328
      |> maps (fn Assumes asms => asms | _ => [])
18795
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1329
      |> map (apsnd (map fst))
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1330
  end;
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1331
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1332
in
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1333
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1334
fun parameters_of thy name =
19278
4d762b77d319 Internal restructuring: local parameters.
ballarin
parents: 19276
diff changeset
  1335
  the_locale thy name |> #params;
18795
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1336
19276
ac90764bb654 New interface function parameters_of_expr.
ballarin
parents: 19061
diff changeset
  1337
fun parameters_of_expr thy expr =
ac90764bb654 New interface function parameters_of_expr.
ballarin
parents: 19061
diff changeset
  1338
  let
ac90764bb654 New interface function parameters_of_expr.
ballarin
parents: 19061
diff changeset
  1339
    val ctxt = ProofContext.init thy;
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1340
    val pts = params_of_expr ctxt [] (intern_expr thy expr)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1341
        ([], Symtab.empty, Symtab.empty);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1342
    val raw_params_elemss = make_raw_params_elemss pts;
19276
ac90764bb654 New interface function parameters_of_expr.
ballarin
parents: 19061
diff changeset
  1343
    val ((_, syn), raw_elemss) = flatten (ctxt, intern_expr thy)
ac90764bb654 New interface function parameters_of_expr.
ballarin
parents: 19061
diff changeset
  1344
        (([], Symtab.empty), Expr expr);
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1345
    val ((parms, _, _), _) =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1346
        read_elemss false ctxt [] (raw_params_elemss @ raw_elemss) [];
19276
ac90764bb654 New interface function parameters_of_expr.
ballarin
parents: 19061
diff changeset
  1347
  in map (fn p as (n, _) => (p, Symtab.lookup syn n |> the)) parms end;
ac90764bb654 New interface function parameters_of_expr.
ballarin
parents: 19061
diff changeset
  1348
18795
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1349
fun local_asms_of thy name =
18890
34a9e466201f *_asms_of fixed.
ballarin
parents: 18831
diff changeset
  1350
  gen_asms_of (single o Library.last_elem) thy name;
18795
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1351
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1352
fun global_asms_of thy name =
18890
34a9e466201f *_asms_of fixed.
ballarin
parents: 18831
diff changeset
  1353
  gen_asms_of I thy name;
18795
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1354
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1355
end;
18795
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1356
303793f49b0f Interface to access the specification of a named locale.
ballarin
parents: 18782
diff changeset
  1357
12546
wenzelm
parents: 12532
diff changeset
  1358
(* full context statements: import + elements + conclusion *)
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1359
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1360
local
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1361
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1362
fun prep_context_statement prep_expr prep_elemss prep_facts
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1363
    do_close fixed_params import elements raw_concl context =
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1364
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1365
    val thy = ProofContext.theory_of context;
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1366
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1367
    val (import_params, import_tenv, import_syn) =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1368
      params_of_expr context fixed_params (prep_expr thy import)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1369
        ([], Symtab.empty, Symtab.empty);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1370
    val includes = map_filter (fn Expr e => SOME e | Elem _ => NONE) elements;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1371
    val (incl_params, incl_tenv, incl_syn) = fold (params_of_expr context fixed_params)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1372
      (map (prep_expr thy) includes) (import_params, import_tenv, import_syn);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1373
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1374
    val ((import_ids, _), raw_import_elemss) =
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1375
      flatten (context, prep_expr thy) (([], Symtab.empty), Expr import);
14215
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
  1376
    (* CB: normalise "includes" among elements *)
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1377
    val ((ids, syn), raw_elemsss) = foldl_map (flatten (context, prep_expr thy))
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1378
      ((import_ids, incl_syn), elements);
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1379
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1380
    val raw_elemss = flat raw_elemsss;
14508
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1381
    (* CB: raw_import_elemss @ raw_elemss is the normalised list of
859b11514537 Experimental command for instantiation of locales in proof contexts:
ballarin
parents: 14446
diff changeset
  1382
       context elements obtained from import and elements. *)
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1383
    (* Now additional elements for parameters are inserted. *)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1384
    val import_params_ids = make_params_ids import_params;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1385
    val incl_params_ids =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1386
        make_params_ids (incl_params \\ import_params);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1387
    val raw_import_params_elemss =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1388
        make_raw_params_elemss (import_params, incl_tenv, incl_syn);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1389
    val raw_incl_params_elemss =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1390
        make_raw_params_elemss (incl_params \\ import_params, incl_tenv, incl_syn);
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1391
    val ((parms, all_elemss, concl), (spec, (_, _, defs))) = prep_elemss do_close
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1392
      context fixed_params
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1393
      (raw_import_params_elemss @ raw_import_elemss @ raw_incl_params_elemss @ raw_elemss) raw_concl;
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1394
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1395
    (* replace extended ids (for axioms) by ids *)
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1396
    val (import_ids', incl_ids) = chop (length import_ids) ids;
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
  1397
    val all_ids = import_params_ids @ import_ids' @ incl_params_ids @ incl_ids;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1398
    val all_elemss' = map (fn (((_, ps), _), (((n, ps'), mode), elems)) =>
17485
c39871c52977 introduced AList module
haftmann
parents: 17449
diff changeset
  1399
        (((n, map (fn p => (p, (the o AList.lookup (op =) ps') p)) ps), mode), elems))
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
  1400
      (all_ids ~~ all_elemss);
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1401
    (* CB: all_elemss and parms contain the correct parameter types *)
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1402
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1403
    val (ps, qs) = chop (length raw_import_params_elemss + length raw_import_elemss) all_elemss';
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1404
    val (import_ctxt, (import_elemss, _)) =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1405
      activate_facts false prep_facts (context, ps);
14215
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
  1406
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1407
    val (ctxt, (elemss, _)) =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1408
      activate_facts false prep_facts (import_ctxt, qs);
12834
e5bec3268932 added locale_facts(_i) interface (useful for simple ML proof tools);
wenzelm
parents: 12806
diff changeset
  1409
  in
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1410
    ((((import_ctxt, import_elemss), (ctxt, elemss, syn)),
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1411
      (parms, spec, defs)), concl)
12834
e5bec3268932 added locale_facts(_i) interface (useful for simple ML proof tools);
wenzelm
parents: 12806
diff changeset
  1412
  end;
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1413
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1414
fun prep_statement prep_locale prep_ctxt raw_locale elems concl ctxt =
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1415
  let
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1416
    val thy = ProofContext.theory_of ctxt;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1417
    val locale = Option.map (prep_locale thy) raw_locale;
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1418
    val (fixed_params, import) =
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1419
      (case locale of
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1420
        NONE => ([], empty)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
  1421
      | SOME name =>
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1422
          let val {params = ps, ...} = the_locale thy name
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1423
          in (map fst ps, Locale name) end);
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1424
    val ((((locale_ctxt, _), (elems_ctxt, _, _)), _), concl') =
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1425
      prep_ctxt false fixed_params import elems concl ctxt;
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1426
  in (locale, locale_ctxt, elems_ctxt, concl') end;
13399
c136276dc847 support locale ``views'' (for cumulative predicates);
wenzelm
parents: 13394
diff changeset
  1427
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1428
fun prep_expr prep import body ctxt =
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1429
  let
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1430
    val (((_, import_elemss), (ctxt', elemss, _)), _) = prep import body ctxt;
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1431
    val all_elems = maps snd (import_elemss @ elemss);
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1432
  in (all_elems, ctxt') end;
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1433
12529
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1434
in
d99716a53f59 simultaneous type-inference of complete context/statement specifications;
wenzelm
parents: 12514
diff changeset
  1435
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1436
val read_ctxt = prep_context_statement intern_expr read_elemss read_facts;
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1437
val cert_ctxt = prep_context_statement (K I) cert_elemss cert_facts;
14215
ebf291f3b449 Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents: 13629
diff changeset
  1438
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1439
fun read_context import body ctxt = #1 (read_ctxt true [] import (map Elem body) [] ctxt);
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1440
fun cert_context import body ctxt = #1 (cert_ctxt true [] import (map Elem body) [] ctxt);
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
  1441
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1442
val read_expr = prep_expr read_context;
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1443
val cert_expr = prep_expr cert_context;
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1444
19932
63bd0eeb4e0d fixed sml/nj value restriction problem
haftmann
parents: 19931
diff changeset
  1445
fun read_context_statement raw_locale = prep_statement intern read_ctxt raw_locale ;
63bd0eeb4e0d fixed sml/nj value restriction problem
haftmann
parents: 19931
diff changeset
  1446
fun cert_context_statement raw_locale = prep_statement (K I) cert_ctxt raw_locale ;
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1447
12502
9e7f72e25022 beginning support for type instantiation;
wenzelm
parents: 12323
diff changeset
  1448
end;
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
  1449
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
  1450
13336
1bd21b082466 tuned add_thmss;
wenzelm
parents: 13308
diff changeset
  1451
(* print locale *)
12070
c72fe1edc9e7 proper treatment of local syntax;
wenzelm
parents: 12063
diff changeset
  1452
17228
19b460b39dad print_locale omits facts by default
ballarin
parents: 17221
diff changeset
  1453
fun print_locale thy show_facts import body =
12070
c72fe1edc9e7 proper treatment of local syntax;
wenzelm
parents: 12063
diff changeset
  1454
  let
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1455
    val (all_elems, ctxt) = read_expr import body (ProofContext.init thy);
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1456
    val prt_elem = Element.pretty_ctxt ctxt;
12277
2b28d7dd91f5 improved ordering of evaluated elements;
wenzelm
parents: 12273
diff changeset
  1457
  in
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1458
    Pretty.big_list "locale elements:" (all_elems
17316
fc7cc8137b97 fixed printing of locales
ballarin
parents: 17271
diff changeset
  1459
      |> (if show_facts then I else filter (fn Notes _ => false | _ => true))
fc7cc8137b97 fixed printing of locales
ballarin
parents: 17271
diff changeset
  1460
      |> map (Pretty.chunks o prt_elem))
13336
1bd21b082466 tuned add_thmss;
wenzelm
parents: 13308
diff changeset
  1461
    |> Pretty.writeln
12277
2b28d7dd91f5 improved ordering of evaluated elements;
wenzelm
parents: 12273
diff changeset
  1462
  end;
12070
c72fe1edc9e7 proper treatment of local syntax;
wenzelm
parents: 12063
diff changeset
  1463
c72fe1edc9e7 proper treatment of local syntax;
wenzelm
parents: 12063
diff changeset
  1464
12706
05fa6a8a6320 removed add_thmss;
wenzelm
parents: 12702
diff changeset
  1465
16144
e339119f4261 renamed cond_extern to extern;
wenzelm
parents: 16105
diff changeset
  1466
(** store results **)
12702
721b622d8967 add_thmss_hybrid;
wenzelm
parents: 12680
diff changeset
  1467
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1468
(* naming of interpreted theorems *)
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1469
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1470
fun global_note_prefix_i kind prfx args thy =
16144
e339119f4261 renamed cond_extern to extern;
wenzelm
parents: 16105
diff changeset
  1471
  thy
19061
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  1472
  |> Theory.qualified_names
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  1473
  |> Theory.sticky_prefix prfx
16144
e339119f4261 renamed cond_extern to extern;
wenzelm
parents: 16105
diff changeset
  1474
  |> PureThy.note_thmss_i kind args
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
  1475
  ||> Theory.restore_naming thy;
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1476
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1477
fun local_note_prefix_i prfx args ctxt =
16144
e339119f4261 renamed cond_extern to extern;
wenzelm
parents: 16105
diff changeset
  1478
  ctxt
19061
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  1479
  |> ProofContext.qualified_names
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  1480
  |> ProofContext.sticky_prefix prfx
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1481
  |> ProofContext.note_thmss_i args
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1482
  ||> ProofContext.restore_naming ctxt;
16144
e339119f4261 renamed cond_extern to extern;
wenzelm
parents: 16105
diff changeset
  1483
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1484
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1485
(* collect witnesses for global registration;
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1486
   requires parameters and flattened list of (assumed!) identifiers
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1487
   instead of recomputing it from the target *)
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1488
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1489
fun collect_global_witnesses thy parms ids vts = let
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1490
    val ts = map Logic.unvarify vts;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1491
    val (parms, parmTs) = split_list parms;
19810
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
  1492
    val parmvTs = map Logic.varifyT parmTs;
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1493
    val vtinst = fold (Sign.typ_match thy) (parmvTs ~~ map Term.fastype_of ts) Vartab.empty;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1494
    val tinst = Vartab.dest vtinst |> map (fn ((x, 0), (_, T)) => (x, T))
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1495
        |> Symtab.make;
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1496
    (* replace parameter names in ids by instantiations *)
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1497
    val vinst = Symtab.make (parms ~~ vts);
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
  1498
    fun vinst_names ps = map (the o Symtab.lookup vinst) ps;
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1499
    val inst = Symtab.make (parms ~~ ts);
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1500
    val ids' = map (apsnd vinst_names) ids;
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1501
    val wits = maps (snd o the o get_global_registration thy) ids';
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1502
  in ((tinst, inst), wits) end;
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1503
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1504
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1505
(* store instantiations of args for all registered interpretations
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1506
   of the theory *)
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1507
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1508
fun note_thmss_registrations kind target args thy =
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1509
  let
19278
4d762b77d319 Internal restructuring: local parameters.
ballarin
parents: 19276
diff changeset
  1510
    val parms = the_locale thy target |> #params |> map fst;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1511
    val ids = flatten (ProofContext.init thy, intern_expr thy)
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  1512
      (([], Symtab.empty), Expr (Locale target)) |> fst |> fst
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1513
      |> map_filter (fn (id, (_, Assumed _)) => SOME id | _ => NONE)
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1514
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1515
    val regs = get_global_registrations thy target;
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1516
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1517
    (* add args to thy for all registrations *)
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1518
18190
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1519
    fun activate (vts, ((prfx, atts2), _)) thy =
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  1520
      let
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1521
        val (insts, prems) = collect_global_witnesses thy parms ids vts;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1522
        val inst_atts =
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1523
          Args.map_values I (Element.instT_type (#1 insts))
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1524
            (Element.inst_term insts) (Element.inst_thm thy insts);
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1525
        val args' = args |> map (fn ((n, atts), [(ths, [])]) =>
19061
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  1526
            ((n, map (Attrib.attribute_i thy) (map inst_atts atts @ atts2)),
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1527
             [(map (Drule.standard o Element.satisfy_thm prems o
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1528
            Element.inst_thm thy insts) ths, [])]));
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1529
      in global_note_prefix_i kind prfx args' thy |> snd end;
18190
b7748c77562f tuned Pattern.match/unify;
wenzelm
parents: 18137
diff changeset
  1530
  in fold activate regs thy end;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1531
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1532
19662
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
  1533
(* term syntax *)
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1534
19662
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
  1535
fun add_term_syntax loc syn =
20037
d4102c7cf051 Fixed erroneous check-in.
ballarin
parents: 20035
diff changeset
  1536
  syn #> ProofContext.map_theory (change_locale loc
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1537
    (fn (axiom, import, elems, params, lparams, term_syntax, regs, intros) =>
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1538
      (axiom, import, elems, params, lparams, (syn, stamp ()) :: term_syntax, regs, intros)));
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1539
19662
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
  1540
fun init_term_syntax loc ctxt =
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
  1541
  fold_rev (fn (f, _) => fn ctxt' => f ctxt')
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
  1542
    (#term_syntax (the_locale (ProofContext.theory_of ctxt) loc)) ctxt;
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1543
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1544
fun init loc =
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1545
  ProofContext.init
19662
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
  1546
  #> init_term_syntax loc
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1547
  #> (#2 o cert_context_statement (SOME loc) [] []);
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1548
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1549
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1550
(* theory/locale results *)
12958
99f5c4a37b29 added smart_have_thmss (global storage);
wenzelm
parents: 12862
diff changeset
  1551
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1552
fun theory_results kind prefix results ctxt thy =
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1553
  let
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1554
    val thy_ctxt = ProofContext.init thy;
20037
d4102c7cf051 Fixed erroneous check-in.
ballarin
parents: 20035
diff changeset
  1555
    val export = singleton (ProofContext.export_standard ctxt thy_ctxt);
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1556
    val facts = map (fn (name, ths) => ((name, []), [(map export ths, [])])) results;
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1557
  in thy |> PureThy.note_thmss_qualified kind prefix facts end;
12958
99f5c4a37b29 added smart_have_thmss (global storage);
wenzelm
parents: 12862
diff changeset
  1558
99f5c4a37b29 added smart_have_thmss (global storage);
wenzelm
parents: 12862
diff changeset
  1559
local
99f5c4a37b29 added smart_have_thmss (global storage);
wenzelm
parents: 12862
diff changeset
  1560
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1561
fun gen_thmss prep_facts global_results kind loc args ctxt thy =
12706
05fa6a8a6320 removed add_thmss;
wenzelm
parents: 12702
diff changeset
  1562
  let
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1563
    val (ctxt', ([(_, [Notes args'])], facts)) =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1564
      activate_facts true prep_facts (ctxt, [((("", []), Assumed []), [Ext (Notes args)])]);
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1565
    val (facts', thy') =
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1566
      thy
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1567
      |> change_locale loc (fn (axiom, import, elems, params, lparams, term_syntax, regs, intros) =>
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1568
        (axiom, import, elems @ [(Notes args', stamp ())], params, lparams, term_syntax, regs, intros))
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1569
      |> note_thmss_registrations kind loc args'
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1570
      |> global_results (map (#1 o #1) args' ~~ map #2 facts) ctxt;
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1571
  in ((facts, facts'), (ProofContext.transfer thy' ctxt', thy')) end;
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
  1572
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1573
fun gen_note prep_locale prep_facts kind raw_loc args thy =
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1574
  let
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1575
    val loc = prep_locale thy raw_loc;
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1576
    val prefix = Sign.base_name loc;
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1577
  in gen_thmss prep_facts (theory_results kind prefix) kind loc args (init loc thy) thy end;
12706
05fa6a8a6320 removed add_thmss;
wenzelm
parents: 12702
diff changeset
  1578
05fa6a8a6320 removed add_thmss;
wenzelm
parents: 12702
diff changeset
  1579
in
05fa6a8a6320 removed add_thmss;
wenzelm
parents: 12702
diff changeset
  1580
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1581
val note_thmss = gen_note intern read_facts;
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1582
val note_thmss_i = gen_note (K I) cert_facts;
12711
6a9412dd7d24 have_thmss vs. have_thmss_i;
wenzelm
parents: 12706
diff changeset
  1583
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1584
fun add_thmss kind loc args ctxt =
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1585
  gen_thmss cert_facts (theory_results kind "")
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1586
    kind loc args ctxt (ProofContext.theory_of ctxt)
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1587
  ||> #1;
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1588
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1589
fun locale_results kind loc args (ctxt, thy) =
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1590
  thy |> gen_thmss cert_facts (K (K (pair [])))
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1591
    kind loc (map (apsnd Thm.simple_fact) args) ctxt
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1592
  |>> #1;
12702
721b622d8967 add_thmss_hybrid;
wenzelm
parents: 12680
diff changeset
  1593
12706
05fa6a8a6320 removed add_thmss;
wenzelm
parents: 12702
diff changeset
  1594
end;
12063
4c16bdee47d4 added add_locale(_i) and store_thm;
wenzelm
parents: 12058
diff changeset
  1595
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
  1596
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1597
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1598
(** define locales **)
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1599
13336
1bd21b082466 tuned add_thmss;
wenzelm
parents: 13308
diff changeset
  1600
(* predicate text *)
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1601
(* CB: generate locale predicates and delta predicates *)
13336
1bd21b082466 tuned add_thmss;
wenzelm
parents: 13308
diff changeset
  1602
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1603
local
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1604
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1605
(* introN: name of theorems for introduction rules of locale and
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1606
     delta predicates;
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1607
   axiomsN: name of theorem set with destruct rules for locale predicates,
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1608
     also name suffix of delta predicates. *)
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1609
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1610
val introN = "intro";
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1611
val axiomsN = "axioms";
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1612
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1613
fun atomize_spec thy ts =
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1614
  let
18502
24efc1587f9d Logic.mk_conjunction_list;
wenzelm
parents: 18473
diff changeset
  1615
    val t = Logic.mk_conjunction_list ts;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1616
    val body = ObjectLogic.atomize_term thy t;
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1617
    val bodyT = Term.fastype_of body;
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1618
  in
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1619
    if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))
18782
e4d9d718b8bb turned abstract_term into ProofContext.abs_def;
wenzelm
parents: 18768
diff changeset
  1620
    else (body, bodyT, ObjectLogic.atomize_cterm (Thm.cterm_of thy t))
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1621
  end;
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1622
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1623
fun aprop_tr' n c = (c, fn args =>
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1624
  if length args = n then Syntax.const "_aprop" $ Term.list_comb (Syntax.free c, args)
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1625
  else raise Match);
13336
1bd21b082466 tuned add_thmss;
wenzelm
parents: 13308
diff changeset
  1626
15104
ballarin
parents: 15099
diff changeset
  1627
(* CB: define one predicate including its intro rule and axioms
ballarin
parents: 15099
diff changeset
  1628
   - bname: predicate name
ballarin
parents: 15099
diff changeset
  1629
   - parms: locale parameters
ballarin
parents: 15099
diff changeset
  1630
   - defs: thms representing substitutions from defines elements
ballarin
parents: 15099
diff changeset
  1631
   - ts: terms representing locale assumptions (not normalised wrt. defs)
ballarin
parents: 15099
diff changeset
  1632
   - norm_ts: terms representing locale assumptions (normalised wrt. defs)
ballarin
parents: 15099
diff changeset
  1633
   - thy: the theory
ballarin
parents: 15099
diff changeset
  1634
*)
ballarin
parents: 15099
diff changeset
  1635
13420
39fca1e5818a removed unused locale_facts(_i);
wenzelm
parents: 13415
diff changeset
  1636
fun def_pred bname parms defs ts norm_ts thy =
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1637
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1638
    val name = Sign.full_name thy bname;
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1639
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1640
    val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1641
    val env = Term.add_term_free_names (body, []);
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1642
    val xs = filter (fn (x, _) => x mem_string env) parms;
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1643
    val Ts = map #2 xs;
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
  1644
    val extraTs = (Term.term_tfrees body \\ foldr Term.add_typ_tfrees [] Ts)
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1645
      |> Library.sort_wrt #1 |> map TFree;
13399
c136276dc847 support locale ``views'' (for cumulative predicates);
wenzelm
parents: 13394
diff changeset
  1646
    val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
13336
1bd21b082466 tuned add_thmss;
wenzelm
parents: 13308
diff changeset
  1647
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1648
    val args = map Logic.mk_type extraTs @ map Free xs;
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1649
    val head = Term.list_comb (Const (name, predT), args);
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  1650
    val statement = ObjectLogic.ensure_propT thy head;
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1651
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 18343
diff changeset
  1652
    val ([pred_def], defs_thy) =
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1653
      thy
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1654
      |> (if bodyT <> propT then I else
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1655
        Theory.add_trfuns ([], [], map (aprop_tr' (length args)) (NameSpace.accesses' name), []))
18671
621efeed255b uniform handling of fixes;
wenzelm
parents: 18617
diff changeset
  1656
      |> Theory.add_consts_i [(bname, predT, NoSyn)]
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1657
      |> PureThy.add_defs_i false [((Thm.def_name bname, Logic.mk_equals (head, body)), [])];
20059
25935807eb08 Element.prove_witness: context;
wenzelm
parents: 20037
diff changeset
  1658
    val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1659
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1660
    val cert = Thm.cterm_of defs_thy;
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1661
20059
25935807eb08 Element.prove_witness: context;
wenzelm
parents: 20037
diff changeset
  1662
    val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ =>
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1663
      Tactic.rewrite_goals_tac [pred_def] THEN
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1664
      Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN
20059
25935807eb08 Element.prove_witness: context;
wenzelm
parents: 20037
diff changeset
  1665
      Tactic.compose_tac (false, Conjunction.intr_list (map (Thm.assume o cert) norm_ts), 0) 1);
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1666
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1667
    val conjuncts =
19423
51eeee99bd8f use conjunction stuff from conjunction.ML;
wenzelm
parents: 19370
diff changeset
  1668
      (Drule.equal_elim_rule2 OF [body_eq,
17257
0ab67cb765da introduced binding priority 1 for linear combinators etc.
haftmann
parents: 17228
diff changeset
  1669
        Tactic.rewrite_rule [pred_def] (Thm.assume (cert statement))])
19423
51eeee99bd8f use conjunction stuff from conjunction.ML;
wenzelm
parents: 19370
diff changeset
  1670
      |> Conjunction.elim_precise [length ts] |> hd;
17257
0ab67cb765da introduced binding priority 1 for linear combinators etc.
haftmann
parents: 17228
diff changeset
  1671
    val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>
20059
25935807eb08 Element.prove_witness: context;
wenzelm
parents: 20037
diff changeset
  1672
      Element.prove_witness defs_ctxt t
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  1673
       (Tactic.rewrite_goals_tac defs THEN
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1674
        Tactic.compose_tac (false, ax, 0) 1));
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1675
  in ((statement, intro, axioms), defs_thy) end;
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1676
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1677
fun assumes_to_notes (Assumes asms) axms =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1678
    fold_map (fn (a, spec) => fn axs =>
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1679
       let val (ps, qs) = chop (length spec) axs
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1680
       in ((a, [(ps, [])]), qs) end) asms axms
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1681
    |> apfst Notes
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1682
  | assumes_to_notes e axms = (e, axms);
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1683
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1684
(* CB: the following two change only "new" elems, these have identifier ("", _). *)
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1685
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1686
(* turn Assumes into Notes elements *)
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1687
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1688
fun change_assumes_elemss axioms elemss =
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1689
  let
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1690
    fun change (id as ("", _), es) =
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1691
          fold_map assumes_to_notes
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1692
            (map (Element.map_ctxt_values I I (Element.satisfy_thm axioms)) es)
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1693
          #-> (fn es' => pair (id, es'))
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1694
      | change e = pair e;
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1695
  in
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1696
    fst (fold_map change elemss (map Element.conclude_witness axioms))
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1697
  end;
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1698
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1699
(* adjust hyps of Notes elements *)
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1700
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1701
fun change_elemss_hyps axioms elemss =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1702
  let
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1703
    fun change (id as ("", _), es) =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1704
        (id, map (fn e as (Notes _) => Element.map_ctxt_values I I (Element.satisfy_thm axioms) e
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1705
                   | e => e) es)
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1706
      | change e = e;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1707
  in map change elemss end;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1708
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1709
in
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1710
15104
ballarin
parents: 15099
diff changeset
  1711
(* CB: main predicate definition function *)
ballarin
parents: 15099
diff changeset
  1712
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1713
fun define_preds bname (parms, ((exts, exts'), (ints, ints')), defs) elemss thy =
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1714
  let
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1715
    val ((elemss', more_ts), a_elem, a_intro, thy') =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1716
      if null exts then ((elemss, []), [], [], thy)
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1717
      else
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1718
        let
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
  1719
          val aname = if null ints then bname else bname ^ "_" ^ axiomsN;
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1720
          val ((statement, intro, axioms), def_thy) =
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1721
            thy |> def_pred aname parms defs exts exts';
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1722
          val elemss' = change_assumes_elemss axioms elemss;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1723
          val def_thy' = def_thy
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1724
            |> PureThy.note_thmss_qualified "" aname [((introN, []), [([intro], [])])]
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1725
            |> snd
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1726
          val a_elem = [(("", []), [Assumes [((bname ^ "_" ^ axiomsN, []), [(statement, [])])]])];
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1727
        in ((elemss', [statement]), a_elem, [intro], def_thy') end;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1728
    val (predicate, stmt', elemss'', b_intro, thy'') =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1729
      if null ints then (([], []), more_ts, elemss' @ a_elem, [], thy')
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1730
      else
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1731
        let
18550
59b89f625d68 add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents: 18546
diff changeset
  1732
          val ((statement, intro, axioms), def_thy) =
13394
b39347206719 define cumulative predicate view;
wenzelm
parents: 13375
diff changeset
  1733
            thy' |> def_pred bname parms defs (ints @ more_ts) (ints' @ more_ts);
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1734
          val cstatement = Thm.cterm_of def_thy statement;
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1735
          val elemss'' = change_elemss_hyps axioms elemss';
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1736
          val def_thy' =
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
  1737
          def_thy
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1738
          |> PureThy.note_thmss_qualified "" bname
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
  1739
               [((introN, []), [([intro], [])]),
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1740
                ((axiomsN, []), [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1741
          |> snd;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1742
          val b_elem = [(("", []),
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1743
               [Assumes [((bname ^ "_" ^ axiomsN, []), [(statement, [])])]])];
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1744
        in (([cstatement], axioms), [statement], elemss'' @ b_elem, [intro], def_thy') end;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1745
  in (((elemss'', predicate, stmt'), (a_intro, b_intro)), thy'') end;
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1746
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1747
end;
13336
1bd21b082466 tuned add_thmss;
wenzelm
parents: 13308
diff changeset
  1748
1bd21b082466 tuned add_thmss;
wenzelm
parents: 13308
diff changeset
  1749
13297
wenzelm
parents: 13211
diff changeset
  1750
(* add_locale(_i) *)
wenzelm
parents: 13211
diff changeset
  1751
wenzelm
parents: 13211
diff changeset
  1752
local
wenzelm
parents: 13211
diff changeset
  1753
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1754
(* turn Defines into Notes elements, accumulate definition terms *)
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1755
19942
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1756
fun defines_to_notes is_ext thy (Defines defs) defns =
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1757
    let
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1758
      val defs' = map (fn (_, (def, _)) => (("", []), (def, []))) defs
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1759
      val notes = map (fn (a, (def, _)) =>
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1760
        (a, [([assume (cterm_of thy def)], [])])) defs
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1761
    in (if is_ext then SOME (Notes notes) else NONE, defns @ [Defines defs']) end
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1762
  | defines_to_notes _ _ e defns = (SOME e, defns);
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1763
19942
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1764
fun change_defines_elemss thy elemss defns =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1765
  let
19942
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1766
    fun change (id as (n, _), es) defns =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1767
        let
19942
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1768
          val (es', defns') = fold_map (defines_to_notes (n="") thy) es defns
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1769
        in ((id, map_filter I es'), defns') end
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1770
  in fold_map change elemss defns end;
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1771
18343
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
  1772
fun gen_add_locale prep_ctxt prep_expr
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
  1773
    do_predicate bname raw_import raw_body thy =
13297
wenzelm
parents: 13211
diff changeset
  1774
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1775
    val name = Sign.full_name thy bname;
18343
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
  1776
    val _ = conditional (is_some (get_locale thy name)) (fn () =>
13297
wenzelm
parents: 13211
diff changeset
  1777
      error ("Duplicate definition of locale " ^ quote name));
wenzelm
parents: 13211
diff changeset
  1778
wenzelm
parents: 13211
diff changeset
  1779
    val thy_ctxt = ProofContext.init thy;
17228
19b460b39dad print_locale omits facts by default
ballarin
parents: 17221
diff changeset
  1780
    val (((import_ctxt, import_elemss), (body_ctxt, body_elemss, syn)),
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1781
      text as (parms, ((_, exts'), _), defs)) =
13375
7cbf2dea46d0 proper predicate definitions of locale body;
wenzelm
parents: 13336
diff changeset
  1782
      prep_ctxt raw_import raw_body thy_ctxt;
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1783
    val elemss = import_elemss @ body_elemss |>
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1784
        map_filter (fn ((id, Assumed axs), elems) => SOME (id, elems) | _ => NONE);
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  1785
    val import = prep_expr thy raw_import;
13297
wenzelm
parents: 13211
diff changeset
  1786
17228
19b460b39dad print_locale omits facts by default
ballarin
parents: 17221
diff changeset
  1787
    val extraTs = foldr Term.add_term_tfrees [] exts' \\
17756
d4a35f82fbb4 minor tweaks for Poplog/ML;
wenzelm
parents: 17496
diff changeset
  1788
      foldr Term.add_typ_tfrees [] (map snd parms);
17228
19b460b39dad print_locale omits facts by default
ballarin
parents: 17221
diff changeset
  1789
    val _ = if null extraTs then ()
17437
9deaf32c83be interpretation uses primitive goal interface
ballarin
parents: 17412
diff changeset
  1790
      else warning ("Additional type variable(s) in locale specification " ^ quote bname);
17228
19b460b39dad print_locale omits facts by default
ballarin
parents: 17221
diff changeset
  1791
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1792
    val ((((elemss', predicate as (predicate_statement, predicate_axioms), stmt'), intros),
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1793
          pred_thy), (import, regs)) =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1794
      if do_predicate then
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1795
        let
19942
dc92e3aebc71 Improved handling of defines imported in duplicate.
ballarin
parents: 19932
diff changeset
  1796
          val (elemss', defns) = change_defines_elemss thy elemss [];
20035
80c79876d2de Internal restructuring: identify no longer computes syntax.
ballarin
parents: 20032
diff changeset
  1797
          val elemss'' = elemss' @ [(("", []), defns)];
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1798
          val (((elemss''', predicate as (_, axioms), stmt'), intros), thy') =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1799
            define_preds bname text elemss'' thy;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1800
          fun mk_regs elemss wits =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1801
            fold_map (fn (id, elems) => fn wts => let
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1802
                val ts = List.concat (List.mapPartial (fn (Assumes asms) =>
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1803
                  SOME (List.concat (map (map #1 o #2) asms)) | _ => NONE) elems);
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1804
                val (wts1, wts2) = chop (length ts) wts
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1805
              in ((apsnd (map fst) id, wts1), wts2) end) elemss wits |> fst;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1806
          val regs = mk_regs elemss''' axioms |>
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1807
                map_filter (fn (("", _), _) => NONE | e => SOME e);
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1808
        in ((((elemss''', predicate, stmt'), intros), thy'), (empty, regs)) end
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1809
      else ((((elemss, ([], []), []), ([], [])), thy), (import, []));
13420
39fca1e5818a removed unused locale_facts(_i);
wenzelm
parents: 13415
diff changeset
  1810
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1811
    fun axiomify axioms elemss =
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1812
      (axioms, elemss) |> foldl_map (fn (axs, (id, elems)) => let
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1813
                   val ts = flat (map_filter (fn (Assumes asms) =>
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1814
                     SOME (maps (map #1 o #2) asms) | _ => NONE) elems);
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1815
                   val (axs1, axs2) = chop (length ts) axs;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  1816
                 in (axs2, ((id, Assumed axs1), elems)) end)
15206
09d78ec709c7 Modified locales: improved implementation of "includes".
ballarin
parents: 15127
diff changeset
  1817
        |> snd;
18569
cf0edebe540c fix: reintroduced pred_ctxt in gen_add_locale
haftmann
parents: 18550
diff changeset
  1818
    val pred_ctxt = ProofContext.init pred_thy;
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1819
    val (ctxt, (_, facts)) = activate_facts true (K I)
18917
77e18862990f refined signature of locale module
haftmann
parents: 18907
diff changeset
  1820
      (pred_ctxt, axiomify predicate_axioms elemss');
20224
9c40a144ee0e moved basic assumption operations from structure ProofContext to Assumption;
wenzelm
parents: 20196
diff changeset
  1821
    val export = singleton (ProofContext.export_standard
9c40a144ee0e moved basic assumption operations from structure ProofContext to Assumption;
wenzelm
parents: 20196
diff changeset
  1822
        (Assumption.add_view thy_ctxt predicate_statement ctxt) thy_ctxt);
13420
39fca1e5818a removed unused locale_facts(_i);
wenzelm
parents: 13415
diff changeset
  1823
    val facts' = facts |> map (fn (a, ths) => ((a, []), [(map export ths, [])]));
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1824
    val elems' = maps #2 (filter (equal "" o #1 o #1) elemss');
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1825
    val elems'' = map_filter (fn (Fixes _) => NONE | e => SOME e) elems';
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1826
    val axs' = map (Element.assume_witness pred_thy) stmt';
19061
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  1827
    val thy' = pred_thy
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  1828
      |> PureThy.note_thmss_qualified "" bname facts' |> snd
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  1829
      |> declare_locale name
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  1830
      |> put_locale name
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1831
       {axiom = axs',
19061
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  1832
        import = import,
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  1833
        elems = map (fn e => (e, stamp ())) elems'',
19278
4d762b77d319 Internal restructuring: local parameters.
ballarin
parents: 19276
diff changeset
  1834
        params = params_of elemss' |> map (fn (x, SOME T) => ((x, T), the (Symtab.lookup syn x))),
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1835
        lparams = map #1 (params_of' body_elemss),
19662
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
  1836
        term_syntax = [],
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1837
        regs = regs,
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1838
        intros = intros};
19293
a67b9916c58e Tuned signature of Locale.add_locale(_i).
ballarin
parents: 19282
diff changeset
  1839
  in ((name, ProofContext.transfer thy' body_ctxt), thy') end;
13297
wenzelm
parents: 13211
diff changeset
  1840
wenzelm
parents: 13211
diff changeset
  1841
in
wenzelm
parents: 13211
diff changeset
  1842
18917
77e18862990f refined signature of locale module
haftmann
parents: 18907
diff changeset
  1843
val add_locale = gen_add_locale read_context intern_expr;
77e18862990f refined signature of locale module
haftmann
parents: 18907
diff changeset
  1844
val add_locale_i = gen_add_locale cert_context (K I);
13297
wenzelm
parents: 13211
diff changeset
  1845
wenzelm
parents: 13211
diff changeset
  1846
end;
wenzelm
parents: 13211
diff changeset
  1847
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  1848
val _ = Context.add_setup
20092
16e5b2723335 Name.internal;
wenzelm
parents: 20069
diff changeset
  1849
 (add_locale_i true "var" empty [Fixes [(Name.internal "x", NONE, NoSyn)]] #> snd #>
16e5b2723335 Name.internal;
wenzelm
parents: 20069
diff changeset
  1850
  add_locale_i true "struct" empty [Fixes [(Name.internal "S", NONE, Structure)]] #> snd);
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
  1851
13297
wenzelm
parents: 13211
diff changeset
  1852
12730
fd0f8fa2b6bd produce predicate text;
wenzelm
parents: 12727
diff changeset
  1853
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1854
(** locale goals **)
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1855
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1856
local
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1857
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1858
fun intern_attrib thy = map_elem (Element.map_ctxt
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1859
  {name = I, var = I, typ = I, term = I, fact = I, attrib = Attrib.intern_src thy});
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1860
18907
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1861
val global_goal = Proof.global_goal ProofDisplay.present_results
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1862
  Attrib.attribute_i ProofContext.bind_propp_schematic_i;
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1863
18907
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1864
fun conclusion prep_att (Element.Shows concl) =
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1865
      (([], concl), fn stmt => fn ctxt => ((Attrib.map_specs prep_att stmt, []), ctxt))
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1866
  | conclusion _ (Element.Obtains cases) = apfst (apfst (map Elem)) (Obtain.statement cases);
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1867
18907
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1868
fun gen_theorem prep_src prep_elem prep_stmt
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1869
    kind before_qed after_qed (name, raw_atts) raw_elems raw_concl thy =
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1870
  let
18907
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1871
    val atts = map (prep_src thy) raw_atts;
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1872
    val ((concl_elems, concl), mk_stmt) = conclusion (prep_src thy) raw_concl;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1873
    val thy_ctxt = ProofContext.init thy;
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1874
    val elems = map (prep_elem thy) (raw_elems @ concl_elems);
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1875
    val (_, _, ctxt, propp) = prep_stmt NONE elems (map snd concl) thy_ctxt;
20037
d4102c7cf051 Fixed erroneous check-in.
ballarin
parents: 20035
diff changeset
  1876
    val ((stmt, facts), goal_ctxt) = mk_stmt (map fst concl ~~ propp) ctxt;
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1877
  in
18907
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1878
    global_goal kind before_qed after_qed (SOME "") (name, atts) stmt goal_ctxt
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1879
    |> Proof.refine_insert facts
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1880
  end;
17449
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  1881
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  1882
fun gen_theorem_in_locale prep_locale prep_src prep_elem prep_stmt no_target
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1883
    kind before_qed after_qed raw_loc (name, atts) raw_elems raw_concl thy =
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1884
  let
18907
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1885
    val ((concl_elems, concl), mk_stmt) = conclusion (prep_src thy) raw_concl;
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1886
    val loc = prep_locale thy raw_loc;
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1887
    val loc_atts = map (prep_src thy) atts;
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1888
    val loc_attss = map (map (prep_src thy) o snd o fst) concl;
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1889
    val target = if no_target then NONE else SOME (extern thy loc);
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1890
    val elems = map (prep_elem thy) (raw_elems @ concl_elems);
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1891
    val names = map (fst o fst) concl;
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1892
19662
2f5d076fde32 replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents: 19585
diff changeset
  1893
    val thy_ctxt = init_term_syntax loc (ProofContext.init thy);
19991
0c9650664d47 Locales no longer generate views. The following functions have changed
ballarin
parents: 19984
diff changeset
  1894
    val (_, loc_ctxt, elems_ctxt, propp) =
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1895
      prep_stmt (SOME raw_loc) elems (map snd concl) thy_ctxt;
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1896
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1897
    val ((stmt, facts), goal_ctxt) =
20037
d4102c7cf051 Fixed erroneous check-in.
ballarin
parents: 20035
diff changeset
  1898
      mk_stmt (map (apsnd (K []) o fst) concl ~~ propp) elems_ctxt;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1899
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  1900
    fun after_qed' results =
19914
fde2b5c0b42b Variable.importT_inst;
wenzelm
parents: 19900
diff changeset
  1901
      let val loc_results = results |> map
20037
d4102c7cf051 Fixed erroneous check-in.
ballarin
parents: 20035
diff changeset
  1902
          (ProofContext.export_standard goal_ctxt loc_ctxt) in
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1903
        curry (locale_results kind loc ((names ~~ loc_attss) ~~ loc_results)) loc_ctxt
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  1904
        #-> (fn res =>
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1905
          if name = "" andalso null loc_atts then I
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  1906
          else #2 o locale_results kind loc [((name, loc_atts), maps #2 res)])
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1907
        #> #2
18806
0d2aa1cceb7d init: include view;
wenzelm
parents: 18795
diff changeset
  1908
        #> after_qed loc_results results
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1909
      end;
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1910
  in
18907
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1911
    global_goal kind before_qed after_qed' target (name, []) stmt goal_ctxt
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1912
    |> Proof.refine_insert facts
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1913
  end;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1914
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1915
in
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1916
18907
f984f22f1cb4 Proof.refine_insert;
wenzelm
parents: 18899
diff changeset
  1917
val theorem = gen_theorem Attrib.intern_src intern_attrib read_context_statement;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1918
val theorem_i = gen_theorem (K I) (K I) cert_context_statement;
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1919
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1920
val theorem_in_locale = gen_theorem_in_locale intern Attrib.intern_src intern_attrib
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1921
  read_context_statement false;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1922
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1923
val theorem_in_locale_i = gen_theorem_in_locale (K I) (K I) (K I)
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1924
  cert_context_statement false;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1925
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1926
val theorem_in_locale_no_target = gen_theorem_in_locale (K I) (K I) (K I)
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  1927
  cert_context_statement true;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1928
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  1929
fun smart_theorem kind NONE a [] (Element.Shows concl) =
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  1930
      Proof.theorem kind NONE (K I) (SOME "") a concl o ProofContext.init
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1931
  | smart_theorem kind NONE a elems concl =
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  1932
      theorem kind NONE (K I) a elems concl
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1933
  | smart_theorem kind (SOME loc) a elems concl =
17856
0551978bfda5 goal statements: accomodate before_qed argument;
wenzelm
parents: 17756
diff changeset
  1934
      theorem_in_locale kind NONE (K (K I)) loc a elems concl;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1935
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1936
end;
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1937
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  1938
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1939
(** Normalisation of locale statements ---
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1940
    discharges goals implied by interpretations **)
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1941
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1942
local
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1943
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1944
fun locale_assm_intros thy =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1945
  Symtab.fold (fn (_, {intros = (a, _), ...}) => fn intros => (a @ intros))
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1946
    (#2 (GlobalLocalesData.get thy)) [];
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1947
fun locale_base_intros thy =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1948
  Symtab.fold (fn (_, {intros = (_, b), ...}) => fn intros => (b @ intros))
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1949
    (#2 (GlobalLocalesData.get thy)) [];
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1950
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1951
fun all_witnesses ctxt =
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1952
  let
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1953
    val thy = ProofContext.theory_of ctxt;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1954
    fun get registrations = Symtab.fold (fn (_, regs) => fn thms =>
20069
77a6b62418bb Witness theorems of interpretations now transfered to current theory.
ballarin
parents: 20059
diff changeset
  1955
        (Registrations.dest thy regs |> map (fn (_, (_, wits)) =>
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1956
          map Element.conclude_witness wits) |> flat) @ thms)
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1957
      registrations [];
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1958
    val globals = get (#3 (GlobalLocalesData.get thy));
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1959
    val locals = get (LocalLocalesData.get ctxt);
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1960
  in globals @ locals end;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1961
(* FIXME: proper varification *)
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1962
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1963
in
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1964
19984
29bb4659f80a Method intro_locales replaced by intro_locales and unfold_locales.
ballarin
parents: 19942
diff changeset
  1965
fun intro_locales_tac eager ctxt facts st =
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1966
  let
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1967
    val wits = all_witnesses ctxt |> map Thm.varifyT;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1968
    val thy = ProofContext.theory_of ctxt;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1969
    val intros = locale_base_intros thy @ (if eager then locale_assm_intros thy else []);
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1970
  in
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1971
    (ALLGOALS (Method.insert_tac facts THEN'
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1972
        REPEAT_ALL_NEW (resolve_tac (wits @ intros)))
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1973
      THEN Tactic.distinct_subgoals_tac) st
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1974
  end;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1975
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1976
val _ = Context.add_setup (Method.add_methods
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1977
  [("intro_locales",
19984
29bb4659f80a Method intro_locales replaced by intro_locales and unfold_locales.
ballarin
parents: 19942
diff changeset
  1978
    Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac false ctxt)),
29bb4659f80a Method intro_locales replaced by intro_locales and unfold_locales.
ballarin
parents: 19942
diff changeset
  1979
    "back-chain introduction rules of locales without unfolding predicates"),
29bb4659f80a Method intro_locales replaced by intro_locales and unfold_locales.
ballarin
parents: 19942
diff changeset
  1980
   ("unfold_locales",
29bb4659f80a Method intro_locales replaced by intro_locales and unfold_locales.
ballarin
parents: 19942
diff changeset
  1981
    Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac true ctxt)),
29bb4659f80a Method intro_locales replaced by intro_locales and unfold_locales.
ballarin
parents: 19942
diff changeset
  1982
    "back-chain all introduction rules of locales")]);
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1983
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1984
end;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  1985
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1986
15598
4ab52355bb53 Registrations of global locale interpretations: improved, better naming.
ballarin
parents: 15596
diff changeset
  1987
(** Interpretation commands **)
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1988
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1989
local
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1990
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1991
(* extract proof obligations (assms and defs) from elements *)
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1992
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  1993
fun extract_asms_elems ((id, Assumed _), elems) = (id, maps Element.prems_of elems)
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1994
  | extract_asms_elems ((id, Derived _), _) = (id, []);
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1995
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1996
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  1997
(* activate instantiated facts in theory or context *)
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  1998
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  1999
fun gen_activate_facts_elemss get_reg note attrib std put_reg add_wit
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2000
        attn all_elemss new_elemss propss thmss thy_ctxt =
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2001
    let
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2002
      fun activate_elem disch (prfx, atts) (Notes facts) thy_ctxt =
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2003
          let
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2004
            val facts' = facts
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2005
              (* discharge hyps in attributes *)
17109
606c269d1e26 added add_locale_context(_i), which returns the body context for presentation;
wenzelm
parents: 17096
diff changeset
  2006
              |> Attrib.map_facts (attrib thy_ctxt o Args.map_values I I I disch)
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2007
              (* insert interpretation attributes *)
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2008
              |> map (apfst (apsnd (fn a => a @ map (attrib thy_ctxt) atts)))
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2009
              (* discharge hyps *)
19061
ffbbac0261c9 removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19025
diff changeset
  2010
              |> map (apsnd (map (apfst (map disch))));
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2011
          in snd (note prfx facts' thy_ctxt) end
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2012
        | activate_elem _ _ _ thy_ctxt = thy_ctxt;
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2013
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2014
      fun activate_elems disch ((id, _), elems) thy_ctxt =
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2015
          let
18343
e36238ac5359 defines: beta/eta contract lhs;
wenzelm
parents: 18330
diff changeset
  2016
            val ((prfx, atts2), _) = the (get_reg thy_ctxt id)
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2017
                handle Option => sys_error ("Unknown registration of " ^
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2018
                  quote (fst id) ^ " while activating facts.");
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2019
          in
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2020
            fold (activate_elem disch (prfx, atts2)) elems thy_ctxt
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2021
          end;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2022
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2023
      val thy_ctxt' = thy_ctxt
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2024
        (* add registrations *)
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2025
        |> fold (fn ((id, _), _) => put_reg id attn) new_elemss
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2026
        (* add witnesses of Assumed elements *)
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2027
        |> fold (fn (id, thms) => fold (add_wit id) thms) (map fst propss ~~ thmss);
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2028
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  2029
      val prems = flat (map_filter
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2030
            (fn ((id, Assumed _), _) => Option.map snd (get_reg thy_ctxt' id)
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2031
              | ((_, Derived _), _) => NONE) all_elemss);
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2032
      val thy_ctxt'' = thy_ctxt'
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2033
        (* add witnesses of Derived elements *)
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2034
        |> fold (fn (id, thms) => fold (add_wit id o Element.satisfy_witness prems) thms)
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  2035
           (map_filter (fn ((_, Assumed _), _) => NONE
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  2036
              | ((id, Derived thms), _) => SOME (id, thms)) new_elemss)
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2037
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2038
      val disch' = std o Element.satisfy_thm prems;  (* FIXME *)
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2039
    in
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2040
      thy_ctxt''
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2041
        (* add facts to theory or context *)
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2042
        |> fold (activate_elems disch') new_elemss
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2043
    end;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2044
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2045
fun global_activate_facts_elemss x = gen_activate_facts_elemss
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2046
      (fn thy => fn (name, ps) =>
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2047
        get_global_registration thy (name, map Logic.varify ps))
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2048
      (global_note_prefix_i "")
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
  2049
      Attrib.attribute_i Drule.standard
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2050
      (fn (name, ps) => put_global_registration (name, map Logic.varify ps))
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2051
      (fn (n, ps) => add_global_witness (n, map Logic.varify ps) o
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  2052
        Element.map_witness (fn (t, th) => (Logic.legacy_unvarify t, legacy_unvarify th))
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2053
        (* FIXME *)) x;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2054
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2055
fun local_activate_facts_elemss x = gen_activate_facts_elemss
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2056
      get_local_registration
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  2057
      local_note_prefix_i
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
  2058
      (Attrib.attribute_i o ProofContext.theory_of) I
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2059
      put_local_registration
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2060
      add_local_witness x;
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2061
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2062
fun gen_prep_registration mk_ctxt is_local read_terms test_reg activate
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2063
    attn expr insts thy_ctxt =
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2064
  let
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2065
    val ctxt = mk_ctxt thy_ctxt;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  2066
    val thy = ProofContext.theory_of ctxt;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2067
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2068
    val ctxt' = ctxt |> ProofContext.theory_of |> ProofContext.init;
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  2069
    val pts = params_of_expr ctxt' [] (intern_expr thy expr)
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  2070
          ([], Symtab.empty, Symtab.empty);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  2071
    val params_ids = make_params_ids (#1 pts);
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  2072
    val raw_params_elemss = make_raw_params_elemss pts;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  2073
    val ((ids, _), raw_elemss) = flatten (ctxt', intern_expr thy)
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16028
diff changeset
  2074
          (([], Symtab.empty), Expr expr);
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2075
    val ((parms, all_elemss, _), (_, (_, defs, _))) =
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  2076
          read_elemss false ctxt' [] (raw_params_elemss @ raw_elemss) [];
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2077
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2078
    (** compute instantiation **)
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2079
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2080
    (* user input *)
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2081
    val insts = if length parms < length insts
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2082
         then error "More arguments than parameters in instantiation."
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2083
         else insts @ replicate (length parms - length insts) NONE;
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2084
    val (ps, pTs) = split_list parms;
19810
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
  2085
    val pvTs = map Logic.legacy_varifyT pTs;
15598
4ab52355bb53 Registrations of global locale interpretations: improved, better naming.
ballarin
parents: 15596
diff changeset
  2086
4ab52355bb53 Registrations of global locale interpretations: improved, better naming.
ballarin
parents: 15596
diff changeset
  2087
    (* instantiations given by user *)
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  2088
    val given = map_filter (fn (_, (NONE, _)) => NONE
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2089
         | (x, (SOME inst, T)) => SOME (x, (inst, T))) (ps ~~ (insts ~~ pvTs));
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2090
    val (given_ps, given_insts) = split_list given;
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2091
    val tvars = foldr Term.add_typ_tvars [] pvTs;
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2092
    val used = foldr Term.add_typ_varnames [] pvTs;
17485
c39871c52977 introduced AList module
haftmann
parents: 17449
diff changeset
  2093
    fun sorts (a, i) = AList.lookup (op =) tvars (a, i);
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2094
    val (vs, vinst) = read_terms thy_ctxt sorts used given_insts;
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2095
    val vars = foldl Term.add_term_tvar_ixns [] vs \\ map fst tvars;
20196
9a19e4de6e2e renamed add_term_varnames to Term.add_varnames (cf. Term.add_vars etc.);
wenzelm
parents: 20167
diff changeset
  2096
    val vars' = fold Term.add_varnames vs vars;
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2097
    val _ = if null vars' then ()
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2098
         else error ("Illegal schematic variable(s) in instantiation: " ^
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2099
           commas_quote (map Syntax.string_of_vname vars'));
15598
4ab52355bb53 Registrations of global locale interpretations: improved, better naming.
ballarin
parents: 15596
diff changeset
  2100
    (* replace new types (which are TFrees) by ones with new names *)
4ab52355bb53 Registrations of global locale interpretations: improved, better naming.
ballarin
parents: 15596
diff changeset
  2101
    val new_Tnames = foldr Term.add_term_tfree_names [] vs;
20092
16e5b2723335 Name.internal;
wenzelm
parents: 20069
diff changeset
  2102
    val new_Tnames' = Name.invent_list used "'a" (length new_Tnames);
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2103
    val renameT =
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2104
          if is_local then I
19810
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
  2105
          else Logic.legacy_unvarifyT o Term.map_type_tfree (fn (a, s) =>
17485
c39871c52977 introduced AList module
haftmann
parents: 17449
diff changeset
  2106
            TFree ((the o AList.lookup (op =) (new_Tnames ~~ new_Tnames')) a, s));
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2107
    val rename =
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2108
          if is_local then I
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2109
          else Term.map_term_types renameT;
15598
4ab52355bb53 Registrations of global locale interpretations: improved, better naming.
ballarin
parents: 15596
diff changeset
  2110
4ab52355bb53 Registrations of global locale interpretations: improved, better naming.
ballarin
parents: 15596
diff changeset
  2111
    val tinst = Symtab.make (map
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2112
                (fn ((x, 0), T) => (x, T |> renameT)
16850
35e07087aba2 sys_error;
wenzelm
parents: 16790
diff changeset
  2113
                  | ((_, n), _) => sys_error "Internal error var in prep_registration") vinst);
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2114
    val inst = Symtab.make (given_ps ~~ map rename vs);
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2115
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2116
    (* defined params without user input *)
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  2117
    val not_given = map_filter (fn (x, (NONE, T)) => SOME (x, T)
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2118
         | (_, (SOME _, _)) => NONE) (ps ~~ (insts ~~ pTs));
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2119
    fun add_def (p, pT) inst =
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2120
      let
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2121
        val (t, T) = case find_first (fn (Free (a, _), _) => a = p) defs of
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2122
               NONE => error ("Instance missing for parameter " ^ quote p)
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2123
             | SOME (Free (_, T), t) => (t, T);
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2124
        val d = Element.inst_term (tinst, inst) t;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2125
      in Symtab.update_new (p, d) inst end;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2126
    val inst = fold add_def not_given inst;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2127
    val insts = (tinst, inst);
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2128
    (* Note: insts contain no vars. *)
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2129
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  2130
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2131
    (** compute proof obligations **)
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2132
15598
4ab52355bb53 Registrations of global locale interpretations: improved, better naming.
ballarin
parents: 15596
diff changeset
  2133
    (* restore "small" ids *)
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2134
    val ids' = map (fn ((n, ps), (_, mode)) =>
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19780
diff changeset
  2135
          ((n, map (fn p => Free (p, (the o AList.lookup (op =) parms) p)) ps), mode))
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  2136
        ids;
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  2137
    val (_, all_elemss') = chop (length raw_params_elemss) all_elemss
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2138
    (* instantiate ids and elements *)
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19914
diff changeset
  2139
    val inst_elemss = (ids' ~~ all_elemss') |> map (fn (((n, ps), _), ((_, mode), elems)) =>
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2140
      ((n, map (Element.inst_term insts) ps),
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2141
        map (fn Int e => Element.inst_ctxt thy insts e) elems)
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2142
      |> apfst (fn id => (id, map_mode (map (Element.inst_witness thy insts)) mode)));
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2143
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2144
    (* remove fragments already registered with theory or context *)
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  2145
    val new_inst_elemss = filter (fn ((id, _), _) =>
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2146
          not (test_reg thy_ctxt id)) inst_elemss;
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
  2147
    val new_ids = map #1 new_inst_elemss;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2148
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2149
    val propss = map extract_asms_elems new_inst_elemss;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2150
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
  2151
    val bind_attrib = Attrib.crude_closure ctxt o Args.assignable;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  2152
    val attn' = apsnd (map (bind_attrib o Attrib.intern_src thy)) attn;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2153
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2154
  in (propss, activate attn' inst_elemss new_inst_elemss propss) end;
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2155
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2156
val prep_global_registration = gen_prep_registration
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2157
     ProofContext.init false
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2158
     (fn thy => fn sorts => fn used =>
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16346
diff changeset
  2159
       Sign.read_def_terms (thy, K NONE, sorts) used true)
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2160
     (fn thy => fn (name, ps) =>
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2161
       test_global_registration thy (name, map Logic.varify ps))
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2162
     global_activate_facts_elemss;
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2163
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2164
val prep_local_registration = gen_prep_registration
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
  2165
     I true
15624
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2166
     (fn ctxt => ProofContext.read_termTs ctxt (K false) (K NONE))
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2167
     smart_test_registration
484178635bd8 Further work on interpretation commands. New command `interpret' for
ballarin
parents: 15623
diff changeset
  2168
     local_activate_facts_elemss;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2169
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2170
fun prep_registration_in_locale target expr thy =
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2171
  (* target already in internal form *)
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2172
  let
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2173
    val ctxt = ProofContext.init thy;
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2174
    val ((raw_target_ids, target_syn), _) = flatten (ctxt, I)
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2175
        (([], Symtab.empty), Expr (Locale target));
19278
4d762b77d319 Internal restructuring: local parameters.
ballarin
parents: 19276
diff changeset
  2176
    val fixed = the_locale thy target |> #params |> map #1;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2177
    val ((all_ids, syn), raw_elemss) = flatten (ctxt, intern_expr thy)
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2178
        ((raw_target_ids, target_syn), Expr expr);
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  2179
    val (target_ids, ids) = chop (length raw_target_ids) all_ids;
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2180
    val ((parms, elemss, _), _) = read_elemss false ctxt fixed raw_elemss [];
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2181
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2182
    (** compute proof obligations **)
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2183
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2184
    (* restore "small" ids, with mode *)
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2185
    val ids' = map (apsnd snd) ids;
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2186
    (* remove Int markers *)
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2187
    val elemss' = map (fn (_, es) =>
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2188
        map (fn Int e => e) es) elemss
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2189
    (* extract assumptions and defs *)
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2190
    val ids_elemss = ids' ~~ elemss';
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2191
    val propss = map extract_asms_elems ids_elemss;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2192
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2193
    (** activation function:
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2194
        - add registrations to the target locale
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2195
        - add induced registrations for all global registrations of
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2196
          the target, unless already present
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2197
        - add facts of induced registrations to theory **)
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2198
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  2199
    val t_ids = map_filter
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2200
        (fn (id, (_, Assumed _)) => SOME id | _ => NONE) target_ids;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2201
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2202
    fun activate thmss thy = let
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2203
        val satisfy = Element.satisfy_thm (flat thmss);
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2204
        val ids_elemss_thmss = ids_elemss ~~ thmss;
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2205
        val regs = get_global_registrations thy target;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2206
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2207
        fun activate_id (((id, Assumed _), _), thms) thy =
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2208
            thy |> put_registration_in_locale target id
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2209
                |> fold (add_witness_in_locale target id) thms
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2210
          | activate_id _ thy = thy;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2211
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2212
        fun activate_reg (vts, ((prfx, atts2), _)) thy = let
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2213
            val (insts, wits) = collect_global_witnesses thy fixed t_ids vts;
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2214
            fun inst_parms ps = map
17485
c39871c52977 introduced AList module
haftmann
parents: 17449
diff changeset
  2215
                  (the o AList.lookup (op =) (map #1 fixed ~~ vts)) ps;
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2216
            val disch = Element.satisfy_thm wits;
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
  2217
            val new_elemss = filter (fn (((name, ps), _), _) =>
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2218
                not (test_global_registration thy (name, inst_parms ps))) (ids_elemss);
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2219
            fun activate_assumed_id (((_, Derived _), _), _) thy = thy
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2220
              | activate_assumed_id ((((name, ps), Assumed _), _), thms) thy = let
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2221
                val ps' = inst_parms ps;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2222
              in
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2223
                if test_global_registration thy (name, ps')
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2224
                then thy
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2225
                else thy
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2226
                  |> put_global_registration (name, ps') (prfx, atts2)
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2227
                  |> fold (fn witn => fn thy => add_global_witness (name, ps')
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2228
                     (Element.inst_witness thy insts witn) thy) thms
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2229
              end;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2230
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2231
            fun activate_derived_id ((_, Assumed _), _) thy = thy
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2232
              | activate_derived_id (((name, ps), Derived ths), _) thy = let
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2233
                val ps' = inst_parms ps;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2234
              in
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2235
                if test_global_registration thy (name, ps')
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2236
                then thy
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2237
                else thy
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2238
                  |> put_global_registration (name, ps') (prfx, atts2)
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2239
                  |> fold (fn witn => fn thy => add_global_witness (name, ps')
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2240
                       (witn |> Element.map_witness (fn (t, th) =>  (* FIXME *)
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2241
                       (Element.inst_term insts t,
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2242
                        disch (Element.inst_thm thy insts (satisfy th))))) thy) ths
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2243
              end;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2244
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2245
            fun activate_elem (Notes facts) thy =
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2246
                let
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2247
                  val facts' = facts
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
  2248
                      |> Attrib.map_facts (Attrib.attribute_i thy o
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2249
                         Args.map_values I (Element.instT_type (#1 insts))
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2250
                           (Element.inst_term insts)
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2251
                           (disch o Element.inst_thm thy insts o satisfy))
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
  2252
                      |> map (apfst (apsnd (fn a => a @ map (Attrib.attribute thy) atts2)))
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2253
                      |> map (apsnd (map (apfst (map (disch o Element.inst_thm thy insts o satisfy)))))
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2254
                in
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
  2255
                  thy
19018
88b4979193d8 added abbreviations: activated by init, no expressions yet;
wenzelm
parents: 18964
diff changeset
  2256
                  |> global_note_prefix_i "" prfx facts'
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
  2257
                  |> snd
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2258
                end
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2259
              | activate_elem _ thy = thy;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2260
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2261
            fun activate_elems (_, elems) thy = fold activate_elem elems thy;
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2262
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2263
          in thy |> fold activate_assumed_id ids_elemss_thmss
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2264
                 |> fold activate_derived_id ids_elemss
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2265
                 |> fold activate_elems new_elemss end;
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2266
      in
17138
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2267
        thy |> fold activate_id ids_elemss_thmss
ad4c98fd367b Interpretation in locales: extended back end;
ballarin
parents: 17109
diff changeset
  2268
            |> fold activate_reg regs
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2269
      end;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2270
17033
f4c1ce91aa3c Release of interpretation in locale.
ballarin
parents: 17000
diff changeset
  2271
  in (propss, activate) end;
17000
552df70f52c2 First version of interpretation in locales. Not yet fully functional.
ballarin
parents: 16947
diff changeset
  2272
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2273
fun prep_propp propss = propss |> map (fn ((name, _), props) =>
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2274
  (("", []), map (rpair [] o Element.mark_witness) props));
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2275
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2276
fun prep_result propps thmss =
19780
dce2168b0ea4 export read/cert_expr;
wenzelm
parents: 19732
diff changeset
  2277
  ListPair.map (fn ((_, props), thms) => map2 Element.make_witness props thms) (propps, thmss);
17437
9deaf32c83be interpretation uses primitive goal interface
ballarin
parents: 17412
diff changeset
  2278
17438
ballarin
parents: 17437
diff changeset
  2279
fun goal_name thy kind target propss =
ballarin
parents: 17437
diff changeset
  2280
    kind ^ Proof.goal_names (Option.map (extern thy) target) ""
ballarin
parents: 17437
diff changeset
  2281
      (propss |> map (fn ((name, _), _) => extern thy name));
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2282
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2283
in
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2284
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2285
fun interpretation (prfx, atts) expr insts thy =
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2286
  let
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2287
    val (propss, activate) = prep_global_registration (prfx, atts) expr insts thy;
17449
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  2288
    val kind = goal_name thy "interpretation" NONE propss;
18502
24efc1587f9d Logic.mk_conjunction_list;
wenzelm
parents: 18473
diff changeset
  2289
    val after_qed = activate o prep_result propss;
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2290
  in
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2291
    ProofContext.init thy
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2292
    |> Proof.theorem_i kind NONE after_qed NONE ("", []) (prep_propp propss)
19810
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
  2293
    |> Element.refine_witness |> Seq.hd
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2294
  end;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2295
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2296
fun interpretation_in_locale (raw_target, expr) thy =
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2297
  let
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2298
    val target = intern thy raw_target;
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2299
    val (propss, activate) = prep_registration_in_locale target expr thy;
17449
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  2300
    val kind = goal_name thy "interpretation" (SOME target) propss;
18137
cb916659c89b moved datatype elem to element.ML;
wenzelm
parents: 18123
diff changeset
  2301
    val after_qed = K (activate o prep_result propss);
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2302
  in
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2303
    thy
18899
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  2304
    |> theorem_in_locale_no_target kind NONE after_qed target ("", []) []
a8e913c93578 theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents: 18890
diff changeset
  2305
      (Element.Shows (prep_propp propss))
19810
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
  2306
    |> Element.refine_witness |> Seq.hd
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2307
  end;
17449
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  2308
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  2309
fun interpret (prfx, atts) expr insts int state =
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  2310
  let
19296
ad96f1095dca interpret: Proof.assert_forward_or_chain;
wenzelm
parents: 19293
diff changeset
  2311
    val _ = Proof.assert_forward_or_chain state;
17449
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  2312
    val ctxt = Proof.context_of state;
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  2313
    val (propss, activate) = prep_local_registration (prfx, atts) expr insts ctxt;
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  2314
    val kind = goal_name (Proof.theory_of state) "interpret" NONE propss;
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2315
    fun after_qed results =
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2316
      Proof.map_context (K (ctxt |> activate (prep_result propss results)))
17449
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  2317
      #> Proof.put_facts NONE
429ca1e21289 interpretation: use goal commands without target -- no storing of results;
wenzelm
parents: 17438
diff changeset
  2318
      #> Seq.single;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2319
  in
18123
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2320
    state
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2321
    |> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i
1bb572e8cee9 avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents: 18038
diff changeset
  2322
      kind NONE after_qed (prep_propp propss)
19810
dae765e552ce renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 19783
diff changeset
  2323
    |> Element.refine_witness |> Seq.hd
17437
9deaf32c83be interpretation uses primitive goal interface
ballarin
parents: 17412
diff changeset
  2324
  end;
15596
8665d08085df First version of global registration command.
ballarin
parents: 15574
diff changeset
  2325
11896
1ff33f896720 moved locale.ML to Isar/locale.ML;
wenzelm
parents:
diff changeset
  2326
end;
17355
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2327
5b31131c0365 load late, after proof.ML;
wenzelm
parents: 17316
diff changeset
  2328
end;