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