src/Pure/Isar/expression.ML
author ballarin
Wed, 17 Dec 2008 17:53:41 +0100
changeset 29239 0a64c3418347
parent 29217 a1c992fb3184
child 29241 3adc06d6504f
permissions -rw-r--r--
Prevent defines from being checked in interpretation.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     1
(*  Title:      Pure/Isar/expression.ML
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     2
    Author:     Clemens Ballarin, TU Muenchen
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     3
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
     4
New locale development --- experimental.
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     5
*)
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     6
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     7
signature EXPRESSION =
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     8
sig
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
     9
  datatype 'term map = Positional of 'term option list | Named of (string * 'term) list;
29214
76c7fc5ba849 Strict prefixes in locales expressions.
ballarin
parents: 29211
diff changeset
    10
  type 'term expr = (string * ((string * bool) * 'term map)) list;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
    11
  type expression = string expr * (Binding.T * string option * mixfix) list;
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
    12
  type expression_i = term expr * (Binding.T * typ option * mixfix) list;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    13
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
    14
  (* Processing of context statements *)
28879
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    15
  val read_statement: Element.context list -> (string * string list) list list ->
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    16
    Proof.context ->  (term * term list) list list * Proof.context;
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    17
  val cert_statement: Element.context_i list -> (term * term list) list list ->
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    18
    Proof.context -> (term * term list) list list * Proof.context;
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    19
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    20
  (* Declaring locales *)
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
    21
  val add_locale_cmd: string -> bstring -> expression -> Element.context list -> theory ->
29028
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
    22
    (string * (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list) *
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
    23
    Proof.context
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
    24
  val add_locale: string -> bstring -> expression_i -> Element.context_i list -> theory ->
29028
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
    25
    (string * (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list) *
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
    26
    Proof.context
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
    27
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
    28
  (* Interpretation *)
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
    29
  val sublocale_cmd: string -> expression -> theory -> Proof.state;
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
    30
  val sublocale: string -> expression_i -> theory -> Proof.state;
29211
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
    31
  val interpretation_cmd: expression -> (Attrib.binding * string) list -> theory -> Proof.state;
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
    32
  val interpretation: expression_i -> (Attrib.binding * term) list -> theory -> Proof.state;
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
    33
  val interpret_cmd: expression -> bool -> Proof.state -> Proof.state;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
    34
  val interpret: expression_i -> bool -> Proof.state -> Proof.state;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    35
end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    36
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    37
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
    38
structure Expression : EXPRESSION =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    39
struct
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    40
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    41
datatype ctxt = datatype Element.ctxt;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    42
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    43
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    44
(*** Expressions ***)
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    45
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    46
datatype 'term map =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    47
  Positional of 'term option list |
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    48
  Named of (string * 'term) list;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    49
29214
76c7fc5ba849 Strict prefixes in locales expressions.
ballarin
parents: 29211
diff changeset
    50
type 'term expr = (string * ((string * bool) * 'term map)) list;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    51
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
    52
type expression = string expr * (Binding.T * string option * mixfix) list;
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
    53
type expression_i = term expr * (Binding.T * typ option * mixfix) list;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    54
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    55
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
    56
(** Internalise locale names in expr **)
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    57
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
    58
fun intern thy instances =  map (apfst (NewLocale.intern thy)) instances;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    59
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    60
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
    61
(** Parameters of expression.
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    62
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
    63
   Sanity check of instantiations and extraction of implicit parameters.
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
    64
   The latter only occurs iff strict = false.
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
    65
   Positional instantiations are extended to match full length of parameter list
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
    66
   of instantiated locale. **)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
    67
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
    68
fun parameters_of thy strict (expr, fixed) =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    69
  let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    70
    fun reject_dups message xs =
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    71
      let val dups = duplicates (op =) xs
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    72
      in
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    73
        if null dups then () else error (message ^ commas dups)
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    74
      end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    75
29006
abe0f11cfa4e Name.name_of -> Binding.base_name
haftmann
parents: 28994
diff changeset
    76
    fun match_bind (n, b) = (n = Binding.base_name b);
29030
0ea94f540548 Order of implicit parameters in locale expression.
ballarin
parents: 29028
diff changeset
    77
    fun parm_eq ((b1, mx1), (b2, mx2)) =
0ea94f540548 Order of implicit parameters in locale expression.
ballarin
parents: 29028
diff changeset
    78
      (* FIXME: cannot compare bindings for equality, instead check for equal name and syntax *)
0ea94f540548 Order of implicit parameters in locale expression.
ballarin
parents: 29028
diff changeset
    79
      (Binding.base_name b1 = Binding.base_name b2) andalso
0ea94f540548 Order of implicit parameters in locale expression.
ballarin
parents: 29028
diff changeset
    80
      (if mx1 = mx2 then true
0ea94f540548 Order of implicit parameters in locale expression.
ballarin
parents: 29028
diff changeset
    81
      else error ("Conflicting syntax for parameter" ^ quote (Binding.display b1) ^
0ea94f540548 Order of implicit parameters in locale expression.
ballarin
parents: 29028
diff changeset
    82
                    " in expression."));
0ea94f540548 Order of implicit parameters in locale expression.
ballarin
parents: 29028
diff changeset
    83
      
29006
abe0f11cfa4e Name.name_of -> Binding.base_name
haftmann
parents: 28994
diff changeset
    84
    fun bind_eq (b1, b2) = (Binding.base_name b1 = Binding.base_name b2);
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    85
      (* FIXME: cannot compare bindings for equality. *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    86
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    87
    fun params_loc loc =
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
    88
          (NewLocale.params_of thy loc |> map (fn (p, _, mx) => (p, mx)), loc);
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    89
    fun params_inst (expr as (loc, (prfx, Positional insts))) =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    90
          let
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    91
            val (ps, loc') = params_loc loc;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    92
	    val d = length ps - length insts;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    93
	    val insts' =
28879
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    94
	      if d < 0 then error ("More arguments than parameters in instantiation of locale " ^
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    95
                quote (NewLocale.extern thy loc))
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    96
	      else insts @ replicate d NONE;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    97
            val ps' = (ps ~~ insts') |>
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    98
              map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE);
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
    99
          in (ps', (loc', (prfx, Positional insts'))) end
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   100
      | params_inst (expr as (loc, (prfx, Named insts))) =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   101
          let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   102
            val _ = reject_dups "Duplicate instantiation of the following parameter(s): "
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   103
              (map fst insts);
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   104
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   105
            val (ps, loc') = params_loc loc;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   106
            val ps' = fold (fn (p, _) => fn ps =>
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   107
              if AList.defined match_bind ps p then AList.delete match_bind p ps
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   108
              else error (quote p ^" not a parameter of instantiated expression.")) insts ps;
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   109
          in (ps', (loc', (prfx, Named insts))) end;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   110
    fun params_expr is =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   111
          let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   112
            val (is', ps') = fold_map (fn i => fn ps =>
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   113
              let
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   114
                val (ps', i') = params_inst i;
29030
0ea94f540548 Order of implicit parameters in locale expression.
ballarin
parents: 29028
diff changeset
   115
                val ps'' = distinct parm_eq (ps @ ps');
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   116
              in (i', ps'') end) is []
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   117
          in (ps', is') end;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   118
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   119
    val (implicit, expr') = params_expr expr;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   120
29006
abe0f11cfa4e Name.name_of -> Binding.base_name
haftmann
parents: 28994
diff changeset
   121
    val implicit' = map (#1 #> Binding.base_name) implicit;
abe0f11cfa4e Name.name_of -> Binding.base_name
haftmann
parents: 28994
diff changeset
   122
    val fixed' = map (#1 #> Binding.base_name) fixed;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   123
    val _ = reject_dups "Duplicate fixed parameter(s): " fixed';
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   124
    val implicit'' = if strict then []
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   125
      else let val _ = reject_dups
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   126
          "Parameter(s) declared simultaneously in expression and for clause: " (implicit' @ fixed')
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   127
        in map (fn (b, mx) => (b, NONE, mx)) implicit end;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   128
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   129
  in (expr', implicit'' @ fixed) end;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   130
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   131
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   132
(** Read instantiation **)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   133
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   134
(* Parse positional or named instantiation *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   135
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   136
local
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   137
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   138
fun prep_inst parse_term parms (Positional insts) ctxt =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   139
      (insts ~~ parms) |> map (fn
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   140
        (NONE, p) => Syntax.parse_term ctxt p |
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   141
        (SOME t, _) => parse_term ctxt t)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   142
  | prep_inst parse_term parms (Named insts) ctxt =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   143
      parms |> map (fn p => case AList.lookup (op =) insts p of
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   144
        SOME t => parse_term ctxt t |
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   145
        NONE => Syntax.parse_term ctxt p);
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   146
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   147
in
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   148
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   149
fun parse_inst x = prep_inst Syntax.parse_term x;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   150
fun make_inst x = prep_inst (K I) x;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   151
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   152
end;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   153
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   154
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   155
(* Instantiation morphism *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   156
29214
76c7fc5ba849 Strict prefixes in locales expressions.
ballarin
parents: 29211
diff changeset
   157
fun inst_morph (parm_names, parm_types) ((prfx, strict), insts') ctxt =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   158
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   159
    (* parameters *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   160
    val type_parm_names = fold Term.add_tfreesT parm_types [] |> map fst;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   161
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   162
    (* type inference and contexts *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   163
    val parm_types' = map (TypeInfer.paramify_vars o Logic.varifyT) parm_types;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   164
    val type_parms = fold Term.add_tvarsT parm_types' [] |> map (Logic.mk_type o TVar);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   165
    val arg = type_parms @ map2 TypeInfer.constrain parm_types' insts';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   166
    val res = Syntax.check_terms ctxt arg;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   167
    val ctxt' = ctxt |> fold Variable.auto_fixes res;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   168
    
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   169
    (* instantiation *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   170
    val (type_parms'', res') = chop (length type_parms) res;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   171
    val insts'' = (parm_names ~~ res') |> map_filter
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   172
      (fn (inst as (x, Free (y, _))) => if x = y then NONE else SOME inst |
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   173
        inst => SOME inst);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   174
    val instT = Symtab.make (type_parm_names ~~ map Logic.dest_type type_parms'');
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   175
    val inst = Symtab.make insts'';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   176
  in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   177
    (Element.inst_morphism (ProofContext.theory_of ctxt) (instT, inst) $>
29214
76c7fc5ba849 Strict prefixes in locales expressions.
ballarin
parents: 29211
diff changeset
   178
      Morphism.binding_morphism (Binding.add_prefix strict prfx), ctxt')
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   179
  end;
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   180
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   181
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   182
(*** Locale processing ***)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   183
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   184
(** Parsing **)
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   185
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   186
fun parse_elem prep_typ prep_term ctxt elem =
29215
f98862eb0591 Use correct mode when parsing elements and conclusion.
ballarin
parents: 29214
diff changeset
   187
  Element.map_ctxt' {binding = I, var = I, typ = prep_typ ctxt,
f98862eb0591 Use correct mode when parsing elements and conclusion.
ballarin
parents: 29214
diff changeset
   188
  term = prep_term (ProofContext.set_mode ProofContext.mode_schematic ctxt), (* FIXME ?? *)
f98862eb0591 Use correct mode when parsing elements and conclusion.
ballarin
parents: 29214
diff changeset
   189
  pat = prep_term (ProofContext.set_mode ProofContext.mode_pattern ctxt),
f98862eb0591 Use correct mode when parsing elements and conclusion.
ballarin
parents: 29214
diff changeset
   190
  fact = I, attrib = I} elem;
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   191
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   192
fun parse_concl prep_term ctxt concl =
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   193
  (map o map) (fn (t, ps) =>
29215
f98862eb0591 Use correct mode when parsing elements and conclusion.
ballarin
parents: 29214
diff changeset
   194
    (prep_term (ProofContext.set_mode ProofContext.mode_schematic ctxt) t, (* FIXME ?? *)
f98862eb0591 Use correct mode when parsing elements and conclusion.
ballarin
parents: 29214
diff changeset
   195
      map (prep_term (ProofContext.set_mode ProofContext.mode_pattern ctxt)) ps)) concl;
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   196
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   197
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   198
(** Simultaneous type inference: instantiations + elements + conclusion **)
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   199
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   200
local
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   201
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   202
fun mk_type T = (Logic.mk_type T, []);
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   203
fun mk_term t = (t, []);
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   204
fun mk_propp (p, pats) = (Syntax.type_constraint propT p, pats);
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   205
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   206
fun dest_type (T, []) = Logic.dest_type T;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   207
fun dest_term (t, []) = t;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   208
fun dest_propp (p, pats) = (p, pats);
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   209
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   210
fun extract_inst (_, (_, ts)) = map mk_term ts;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   211
fun restore_inst ((l, (p, _)), cs) = (l, (p, map dest_term cs));
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   212
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   213
fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   214
  | extract_elem (Constrains csts) = map (#2 #> single #> map mk_type) csts
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   215
  | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   216
  | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   217
  | extract_elem (Notes _) = [];
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   218
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   219
fun restore_elem (Fixes fixes, css) =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   220
      (fixes ~~ css) |> map (fn ((x, _, mx), cs) =>
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   221
        (x, cs |> map dest_type |> try hd, mx)) |> Fixes
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   222
  | restore_elem (Constrains csts, css) =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   223
      (csts ~~ css) |> map (fn ((x, _), cs) =>
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   224
        (x, cs |> map dest_type |> hd)) |> Constrains
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   225
  | restore_elem (Assumes asms, css) =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   226
      (asms ~~ css) |> map (fn ((b, _), cs) => (b, map dest_propp cs)) |> Assumes
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   227
  | restore_elem (Defines defs, css) =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   228
      (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   229
  | restore_elem (Notes notes, _) = Notes notes;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   230
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   231
fun check cs context =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   232
  let
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   233
    fun prep (_, pats) (ctxt, t :: ts) =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   234
      let val ctxt' = Variable.auto_fixes t ctxt
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   235
      in
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   236
        ((t, Syntax.check_props (ProofContext.set_mode ProofContext.mode_pattern ctxt') pats),
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   237
          (ctxt', ts))
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   238
      end
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   239
    val (cs', (context', _)) = fold_map prep cs
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   240
      (context, Syntax.check_terms
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   241
        (ProofContext.set_mode ProofContext.mode_schematic context) (map fst cs));
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   242
  in (cs', context') end;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   243
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   244
in
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   245
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   246
fun check_autofix insts elems concl ctxt =
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   247
  let
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   248
    val inst_cs = map extract_inst insts;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   249
    val elem_css = map extract_elem elems;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   250
    val concl_cs = (map o map) mk_propp concl;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   251
    (* Type inference *)
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   252
    val (inst_cs' :: css', ctxt') =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   253
      (fold_burrow o fold_burrow) check (inst_cs :: elem_css @ [concl_cs]) ctxt;
28936
f1647bf418f5 No resolution of patterns within context statements.
ballarin
parents: 28903
diff changeset
   254
    val (elem_css', [concl_cs']) = chop (length elem_css) css';
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   255
  in
28936
f1647bf418f5 No resolution of patterns within context statements.
ballarin
parents: 28903
diff changeset
   256
    (map restore_inst (insts ~~ inst_cs'), map restore_elem (elems ~~ elem_css'),
f1647bf418f5 No resolution of patterns within context statements.
ballarin
parents: 28903
diff changeset
   257
      concl_cs', ctxt')
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   258
  end;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   259
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   260
end;
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   261
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   262
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   263
(** Prepare locale elements **)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   264
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   265
fun declare_elem prep_vars (Fixes fixes) ctxt =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   266
      let val (vars, _) = prep_vars fixes ctxt
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   267
      in ctxt |> ProofContext.add_fixes_i vars |> snd end
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   268
  | declare_elem prep_vars (Constrains csts) ctxt =
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   269
      ctxt |> prep_vars (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) csts) |> snd
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   270
  | declare_elem _ (Assumes _) ctxt = ctxt
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   271
  | declare_elem _ (Defines _) ctxt = ctxt
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   272
  | declare_elem _ (Notes _) ctxt = ctxt;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   273
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   274
(** Finish locale elements, extract specification text **)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   275
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   276
val norm_term = Envir.beta_norm oo Term.subst_atomic;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   277
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   278
fun bind_def false ctxt eq (xs, env, eqs) =
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   279
      let val b' = norm_term env eq;
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   280
      in (Term.add_frees b' xs, env, eqs) end
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   281
  | bind_def true ctxt eq (xs, env, eqs) =
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   282
      let
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   283
        val _ = LocalDefs.cert_def ctxt eq;
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   284
        val ((y, T), b) = LocalDefs.abs_def eq;
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   285
        val b' = norm_term env b;
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   286
        fun err msg = error (msg ^ ": " ^ quote y);
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   287
      in
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   288
        exists (fn (x, _) => x = y) xs andalso
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   289
          err "Attempt to define previously specified variable";
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   290
        exists (fn (Free (y', _), _) => y = y' | _ => false) env andalso
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   291
          err "Attempt to redefine variable";
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   292
        (Term.add_frees b' xs, (Free (y, T), b') :: env, eq :: eqs)
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   293
      end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   294
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   295
fun eval_text _ _ _ (Fixes _) text = text
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   296
  | eval_text _ _ _ (Constrains _) text = text
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   297
  | eval_text _ _ is_ext (Assumes asms)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   298
        (((exts, exts'), (ints, ints')), (xs, env, defs)) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   299
      let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   300
        val ts = maps (map #1 o #2) asms;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   301
        val ts' = map (norm_term env) ts;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   302
        val spec' =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   303
          if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   304
          else ((exts, exts'), (ints @ ts, ints' @ ts'));
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   305
      in (spec', (fold Term.add_frees ts' xs, env, defs)) end
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   306
  | eval_text ctxt check _ (Defines defs) (spec, binds) =
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   307
      (spec, fold (bind_def check ctxt o #1 o #2) defs binds)
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   308
  | eval_text _ _ _ (Notes _) text = text;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   309
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   310
fun closeup _ _ false elem = elem
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   311
  | closeup ctxt parms true elem =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   312
      let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   313
        fun close_frees t =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   314
          let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   315
            val rev_frees =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   316
              Term.fold_aterms (fn Free (x, T) =>
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   317
                if AList.defined (op =) parms x then I else insert (op =) (x, T) | _ => I) t [];
29021
ce100fbc3c8e Proper shape of assumptions generated from Defines elements.
ballarin
parents: 29020
diff changeset
   318
          in Term.list_all_free (rev rev_frees, t) end; (* FIXME use fold Logic.all *)
29019
8e7d6f959bd7 Explicitly close up defines.
ballarin
parents: 29018
diff changeset
   319
  (* FIXME consider closing in syntactic phase *)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   320
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   321
        fun no_binds [] = []
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   322
          | no_binds _ = error "Illegal term bindings in context element";
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   323
      in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   324
        (case elem of
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   325
          Assumes asms => Assumes (asms |> map (fn (a, propps) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   326
            (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
29022
54d3a31ca0f6 Default names for definitions.
ballarin
parents: 29021
diff changeset
   327
        | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) =>
54d3a31ca0f6 Default names for definitions.
ballarin
parents: 29021
diff changeset
   328
            let val ((c, _), t') = LocalDefs.cert_def ctxt (close_frees t)
54d3a31ca0f6 Default names for definitions.
ballarin
parents: 29021
diff changeset
   329
            in
54d3a31ca0f6 Default names for definitions.
ballarin
parents: 29021
diff changeset
   330
              ((Binding.map_base (Thm.def_name_optional c) name, atts), (t', no_binds ps))
54d3a31ca0f6 Default names for definitions.
ballarin
parents: 29021
diff changeset
   331
            end))
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   332
        | e => e)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   333
      end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   334
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   335
fun finish_primitive parms _ (Fixes fixes) = Fixes (map (fn (binding, _, mx) =>
29006
abe0f11cfa4e Name.name_of -> Binding.base_name
haftmann
parents: 28994
diff changeset
   336
      let val x = Binding.base_name binding
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   337
      in (binding, AList.lookup (op =) parms x, mx) end) fixes)
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   338
  | finish_primitive _ _ (Constrains _) = Constrains []
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   339
  | finish_primitive _ close (Assumes asms) = close (Assumes asms)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   340
  | finish_primitive _ close (Defines defs) = close (Defines defs)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   341
  | finish_primitive _ _ (Notes facts) = Notes facts;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   342
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   343
fun finish_inst ctxt parms check_defs do_close (loc, (prfx, inst)) text =
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   344
  let
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   345
    val thy = ProofContext.theory_of ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   346
    val (parm_names, parm_types) = NewLocale.params_of thy loc |>
29006
abe0f11cfa4e Name.name_of -> Binding.base_name
haftmann
parents: 28994
diff changeset
   347
      map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   348
    val (asm, defs) = NewLocale.specification_of thy loc;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   349
    val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst) ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   350
    val asm' = Option.map (Morphism.term morph) asm;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   351
    val defs' = map (Morphism.term morph) defs;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   352
    val text' = text |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   353
      (if is_some asm
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   354
        then eval_text ctxt check_defs false (Assumes [(Attrib.empty_binding, [(the asm', [])])])
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   355
        else I) |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   356
      (if not (null defs)
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   357
        then eval_text ctxt check_defs false (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs'))
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   358
        else I)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   359
(* FIXME clone from new_locale.ML *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   360
  in ((loc, morph), text') end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   361
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   362
fun finish_elem ctxt parms check_defs do_close elem text =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   363
  let
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   364
    val elem' = finish_primitive parms (closeup ctxt parms do_close) elem;
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   365
    val text' = eval_text ctxt check_defs true elem' text;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   366
  in (elem', text') end
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   367
  
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   368
fun finish ctxt parms check_defs do_close insts elems text =
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   369
  let
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   370
    val (deps, text') = fold_map (finish_inst ctxt parms check_defs do_close) insts text;
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   371
    val (elems', text'') = fold_map (finish_elem ctxt parms check_defs do_close) elems text';
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   372
  in (deps, elems', text'') end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   373
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   374
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   375
(** Process full context statement: instantiations + elements + conclusion **)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   376
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   377
(* Interleave incremental parsing and type inference over entire parsed stretch. *)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   378
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   379
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   380
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   381
fun prep_full_context_statement parse_typ parse_prop parse_inst prep_vars prep_expr
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   382
    strict do_close context raw_import raw_elems raw_concl =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   383
  let
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   384
    val thy = ProofContext.theory_of context;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   385
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   386
    val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import);
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   387
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   388
    fun prep_inst (loc, (prfx, inst)) (i, insts, ctxt) =
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   389
      let
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   390
        val (parm_names, parm_types) = NewLocale.params_of thy loc |>
29006
abe0f11cfa4e Name.name_of -> Binding.base_name
haftmann
parents: 28994
diff changeset
   391
          map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   392
        val inst' = parse_inst parm_names inst ctxt;
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   393
        val parm_types' = map (TypeInfer.paramify_vars o
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   394
          Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) o Logic.varifyT) parm_types;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   395
        val inst'' = map2 TypeInfer.constrain parm_types' inst';
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   396
        val insts' = insts @ [(loc, (prfx, inst''))];
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   397
        val (insts'', _, _, ctxt' (* FIXME not used *) ) = check_autofix insts' [] [] ctxt;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   398
        val inst''' = insts'' |> List.last |> snd |> snd;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   399
        val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst''') ctxt;
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   400
        val ctxt'' = NewLocale.activate_declarations thy (loc, morph) ctxt;
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   401
      in (i+1, insts', ctxt'') end;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   402
  
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   403
    fun prep_elem raw_elem (insts, elems, ctxt) =
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   404
      let
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   405
        val ctxt' = declare_elem prep_vars raw_elem ctxt;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   406
        val elems' = elems @ [parse_elem parse_typ parse_prop ctxt' raw_elem];
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   407
        val (_, _, _, ctxt'') = check_autofix insts elems' [] ctxt';
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   408
      in (insts, elems', ctxt') end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   409
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   410
    fun prep_concl raw_concl (insts, elems, ctxt) =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   411
      let
29215
f98862eb0591 Use correct mode when parsing elements and conclusion.
ballarin
parents: 29214
diff changeset
   412
        val concl = parse_concl parse_prop ctxt raw_concl;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   413
      in check_autofix insts elems concl ctxt end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   414
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   415
    val fors = prep_vars fixed context |> fst;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   416
    val ctxt = context |> ProofContext.add_fixes_i fors |> snd;
29206
62dc8762ec00 Preserve idents (expression in sublocale).
ballarin
parents: 29035
diff changeset
   417
    val (_, insts', ctxt') = fold prep_inst raw_insts (0, [], ctxt);
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   418
    val (_, elems'', ctxt'') = fold prep_elem raw_elems (insts', [], ctxt');
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   419
    val (insts, elems, concl, ctxt) = prep_concl raw_concl (insts', elems'', ctxt'');
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   420
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   421
    (* Retrieve parameter types *)
29006
abe0f11cfa4e Name.name_of -> Binding.base_name
haftmann
parents: 28994
diff changeset
   422
    val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Binding.base_name o #1) fixes) |
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   423
      _ => fn ps => ps) (Fixes fors :: elems) [];
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   424
    val (Ts, ctxt') = fold_map ProofContext.inferred_param xs ctxt; 
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   425
    val parms = xs ~~ Ts;  (* params from expression and elements *)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   426
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   427
    val Fixes fors' = finish_primitive parms I (Fixes fors);
29239
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   428
    val (deps, elems', text) =
0a64c3418347 Prevent defines from being checked in interpretation.
ballarin
parents: 29217
diff changeset
   429
      finish ctxt' parms (not strict) do_close insts elems ((([], []), ([], [])), ([], [], []));
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   430
    (* text has the following structure:
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   431
           (((exts, exts'), (ints, ints')), (xs, env, defs))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   432
       where
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   433
         exts: external assumptions (terms in assumes elements)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   434
         exts': dito, normalised wrt. env
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   435
         ints: internal assumptions (terms in assumptions from insts)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   436
         ints': dito, normalised wrt. env
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   437
         xs: the free variables in exts' and ints' and rhss of definitions,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   438
           this includes parameters except defined parameters
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   439
         env: list of term pairs encoding substitutions, where the first term
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   440
           is a free variable; substitutions represent defines elements and
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   441
           the rhs is normalised wrt. the previous env
29031
e74341997a48 Pass on defines in inheritance; reject illicit defines created by instantiation.
ballarin
parents: 29030
diff changeset
   442
         defs: the equations from the defines elements
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   443
       elems is an updated version of raw_elems:
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   444
         - type info added to Fixes and modified in Constrains
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   445
         - axiom and definition statement replaced by corresponding one
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   446
           from proppss in Assumes and Defines
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   447
         - Facts unchanged
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   448
       *)
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   449
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   450
  in ((fors', deps, elems', concl), (parms, text)) end
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   451
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   452
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   453
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   454
fun read_full_context_statement x =
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   455
  prep_full_context_statement Syntax.parse_typ Syntax.parse_prop parse_inst
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   456
  ProofContext.read_vars intern x;
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   457
fun cert_full_context_statement x =
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   458
  prep_full_context_statement (K I) (K I) make_inst ProofContext.cert_vars (K I) x;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   459
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   460
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   461
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   462
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   463
(* Context statement: elements + conclusion *)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   464
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   465
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   466
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   467
fun prep_statement prep activate raw_elems raw_concl context =
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   468
  let
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   469
     val ((_, _, elems, concl), _) = prep true false context ([], []) raw_elems raw_concl;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   470
     val (_, context') = activate elems (ProofContext.set_stmt true context);
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   471
  in (concl, context') end;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   472
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   473
in
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   474
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   475
fun read_statement x = prep_statement read_full_context_statement Element.activate x;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   476
fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   477
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   478
end;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   479
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   480
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   481
(* Locale declaration: import + elements *)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   482
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   483
local
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   484
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   485
fun prep_declaration prep activate raw_import raw_elems context =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   486
  let
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   487
    val ((fixed, deps, elems, _), (parms, (spec, (_, _, defs)))) =
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   488
      prep false true context raw_import raw_elems [];
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   489
    (* Declare parameters and imported facts *)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   490
    val context' = context |>
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   491
      ProofContext.add_fixes_i fixed |> snd |>
29206
62dc8762ec00 Preserve idents (expression in sublocale).
ballarin
parents: 29035
diff changeset
   492
      fold NewLocale.activate_local_facts deps;
29217
a1c992fb3184 Finer-grained activation so that facts from earlier elements are available.
ballarin
parents: 29215
diff changeset
   493
    val (elems', _) = activate elems (ProofContext.set_stmt true context');
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   494
  in ((fixed, deps, elems'), (parms, spec, defs)) end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   495
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   496
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   497
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   498
fun read_declaration x = prep_declaration read_full_context_statement Element.activate x;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   499
fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   500
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   501
end;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   502
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   503
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   504
(* Locale expression to set up a goal *)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   505
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   506
local
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   507
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   508
fun props_of thy (name, morph) =
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   509
  let
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   510
    val (asm, defs) = NewLocale.specification_of thy name;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   511
  in
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   512
    (case asm of NONE => defs | SOME asm => asm :: defs) |> map (Morphism.term morph)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   513
  end;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   514
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   515
fun prep_goal_expression prep expression context =
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   516
  let
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   517
    val thy = ProofContext.theory_of context;
28879
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
   518
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   519
    val ((fixed, deps, _, _), _) = prep true true context expression [] [];
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   520
    (* proof obligations *)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   521
    val propss = map (props_of thy) deps;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   522
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   523
    val goal_ctxt = context |>
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   524
      ProofContext.add_fixes_i fixed |> snd |>
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   525
      (fold o fold) Variable.auto_fixes propss;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   526
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   527
    val export = Variable.export_morphism goal_ctxt context;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   528
    val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   529
    val exp_term = Drule.term_rule thy (singleton exp_fact);
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   530
    val exp_typ = Logic.type_map exp_term;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   531
    val export' =
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   532
      Morphism.morphism {binding = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   533
  in ((propss, deps, export'), goal_ctxt) end;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   534
    
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   535
in
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   536
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   537
fun read_goal_expression x = prep_goal_expression read_full_context_statement x;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   538
fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x;
28879
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
   539
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   540
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   541
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   542
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   543
(*** Locale declarations ***)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   544
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   545
(* axiomsN: name of theorem set with destruct rules for locale predicates,
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   546
     also name suffix of delta predicates and assumptions. *)
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   547
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   548
val axiomsN = "axioms";
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   549
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   550
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   551
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   552
(* introN: name of theorems for introduction rules of locale and
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   553
     delta predicates *)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   554
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   555
val introN = "intro";
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   556
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   557
fun atomize_spec thy ts =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   558
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   559
    val t = Logic.mk_conjunction_balanced ts;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   560
    val body = ObjectLogic.atomize_term thy t;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   561
    val bodyT = Term.fastype_of body;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   562
  in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   563
    if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   564
    else (body, bodyT, ObjectLogic.atomize (Thm.cterm_of thy t))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   565
  end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   566
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   567
(* achieve plain syntax for locale predicates (without "PROP") *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   568
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   569
fun aprop_tr' n c = (Syntax.constN ^ c, fn ctxt => fn args =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   570
  if length args = n then
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   571
    Syntax.const "_aprop" $
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   572
      Term.list_comb (Syntax.free (Consts.extern (ProofContext.consts_of ctxt) c), args)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   573
  else raise Match);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   574
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   575
(* define one predicate including its intro rule and axioms
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   576
   - bname: predicate name
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   577
   - parms: locale parameters
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   578
   - defs: thms representing substitutions from defines elements
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   579
   - ts: terms representing locale assumptions (not normalised wrt. defs)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   580
   - norm_ts: terms representing locale assumptions (normalised wrt. defs)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   581
   - thy: the theory
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   582
*)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   583
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   584
fun def_pred bname parms defs ts norm_ts thy =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   585
  let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   586
    val name = Sign.full_bname thy bname;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   587
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   588
    val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   589
    val env = Term.add_term_free_names (body, []);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   590
    val xs = filter (member (op =) env o #1) parms;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   591
    val Ts = map #2 xs;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   592
    val extraTs = (Term.term_tfrees body \\ fold Term.add_tfreesT Ts [])
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   593
      |> Library.sort_wrt #1 |> map TFree;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   594
    val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   595
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   596
    val args = map Logic.mk_type extraTs @ map Free xs;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   597
    val head = Term.list_comb (Const (name, predT), args);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   598
    val statement = ObjectLogic.ensure_propT thy head;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   599
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   600
    val ([pred_def], defs_thy) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   601
      thy
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   602
      |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   603
      |> Sign.declare_const [] ((Binding.name bname, predT), NoSyn) |> snd
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   604
      |> PureThy.add_defs false
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   605
        [((Thm.def_name bname, Logic.mk_equals (head, body)), [Thm.kind_internal])];
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   606
    val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   607
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   608
    val cert = Thm.cterm_of defs_thy;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   609
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   610
    val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   611
      MetaSimplifier.rewrite_goals_tac [pred_def] THEN
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   612
      Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   613
      Tactic.compose_tac (false,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   614
        Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   615
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   616
    val conjuncts =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   617
      (Drule.equal_elim_rule2 OF [body_eq,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   618
        MetaSimplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))])
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   619
      |> Conjunction.elim_balanced (length ts);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   620
    val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   621
      Element.prove_witness defs_ctxt t
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   622
       (MetaSimplifier.rewrite_goals_tac defs THEN
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   623
        Tactic.compose_tac (false, ax, 0) 1));
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   624
  in ((statement, intro, axioms), defs_thy) end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   625
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   626
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   627
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   628
(* CB: main predicate definition function *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   629
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   630
fun define_preds pname (parms, ((exts, exts'), (ints, ints')), defs) thy =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   631
  let
29031
e74341997a48 Pass on defines in inheritance; reject illicit defines created by instantiation.
ballarin
parents: 29030
diff changeset
   632
    val defs' = map (cterm_of thy #> Assumption.assume #> Drule.gen_all #> Drule.abs_def) defs;
e74341997a48 Pass on defines in inheritance; reject illicit defines created by instantiation.
ballarin
parents: 29030
diff changeset
   633
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   634
    val (a_pred, a_intro, a_axioms, thy'') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   635
      if null exts then (NONE, NONE, [], thy)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   636
      else
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   637
        let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   638
          val aname = if null ints then pname else pname ^ "_" ^ axiomsN;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   639
          val ((statement, intro, axioms), thy') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   640
            thy
29031
e74341997a48 Pass on defines in inheritance; reject illicit defines created by instantiation.
ballarin
parents: 29030
diff changeset
   641
            |> def_pred aname parms defs' exts exts';
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   642
          val (_, thy'') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   643
            thy'
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   644
            |> Sign.add_path aname
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   645
            |> Sign.no_base_names
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   646
            |> PureThy.note_thmss Thm.internalK
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   647
              [((Binding.name introN, []), [([intro], [NewLocale.unfold_attrib])])]
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   648
            ||> Sign.restore_naming thy';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   649
          in (SOME statement, SOME intro, axioms, thy'') end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   650
    val (b_pred, b_intro, b_axioms, thy'''') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   651
      if null ints then (NONE, NONE, [], thy'')
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   652
      else
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   653
        let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   654
          val ((statement, intro, axioms), thy''') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   655
            thy''
29031
e74341997a48 Pass on defines in inheritance; reject illicit defines created by instantiation.
ballarin
parents: 29030
diff changeset
   656
            |> def_pred pname parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred);
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   657
          val (_, thy'''') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   658
            thy'''
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   659
            |> Sign.add_path pname
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   660
            |> Sign.no_base_names
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   661
            |> PureThy.note_thmss Thm.internalK
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   662
                 [((Binding.name introN, []), [([intro], [NewLocale.intro_attrib])]),
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   663
                  ((Binding.name axiomsN, []),
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   664
                    [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   665
            ||> Sign.restore_naming thy''';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   666
        in (SOME statement, SOME intro, axioms, thy'''') end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   667
  in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy'''') end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   668
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   669
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   670
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   671
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   672
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   673
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   674
fun assumes_to_notes (Assumes asms) axms =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   675
      fold_map (fn (a, spec) => fn axs =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   676
          let val (ps, qs) = chop (length spec) axs
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   677
          in ((a, [(ps, [])]), qs) end) asms axms
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   678
      |> apfst (curry Notes Thm.assumptionK)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   679
  | assumes_to_notes e axms = (e, axms);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   680
29031
e74341997a48 Pass on defines in inheritance; reject illicit defines created by instantiation.
ballarin
parents: 29030
diff changeset
   681
fun defines_to_notes thy (Defines defs) =
e74341997a48 Pass on defines in inheritance; reject illicit defines created by instantiation.
ballarin
parents: 29030
diff changeset
   682
      Notes (Thm.definitionK, map (fn (a, (def, _)) =>
e74341997a48 Pass on defines in inheritance; reject illicit defines created by instantiation.
ballarin
parents: 29030
diff changeset
   683
        (a, [([Assumption.assume (cterm_of thy def)], [])])) defs)
e74341997a48 Pass on defines in inheritance; reject illicit defines created by instantiation.
ballarin
parents: 29030
diff changeset
   684
  | defines_to_notes _ e = e;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   685
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   686
fun gen_add_locale prep_decl
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   687
    bname predicate_name raw_imprt raw_body thy =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   688
  let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   689
    val name = Sign.full_bname thy bname;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   690
    val _ = NewLocale.test_locale thy name andalso
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   691
      error ("Duplicate definition of locale " ^ quote name);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   692
29031
e74341997a48 Pass on defines in inheritance; reject illicit defines created by instantiation.
ballarin
parents: 29030
diff changeset
   693
    val ((fixed, deps, body_elems), text as (parms, ((_, exts'), _), defs)) =
29021
ce100fbc3c8e Proper shape of assumptions generated from Defines elements.
ballarin
parents: 29020
diff changeset
   694
      prep_decl raw_imprt raw_body (ProofContext.init thy);
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   695
    val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   696
      define_preds predicate_name text thy;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   697
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   698
    val extraTs = fold Term.add_tfrees exts' [] \\ fold Term.add_tfreesT (map snd parms) [];
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   699
    val _ = if null extraTs then ()
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   700
      else warning ("Additional type variable(s) in locale specification " ^ quote bname);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   701
29035
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   702
    val a_satisfy = Element.satisfy_morphism a_axioms;
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   703
    val b_satisfy = Element.satisfy_morphism b_axioms;
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   704
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   705
    val params = fixed @
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   706
      (body_elems |> map_filter (fn Fixes fixes => SOME fixes | _ => NONE) |> flat);
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   707
    val asm = if is_some b_statement then b_statement else a_statement;
29028
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
   708
29035
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   709
    (* These are added immediately. *)
29028
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
   710
    val notes =
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
   711
        if is_some asm
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
   712
        then [(Thm.internalK, [((Binding.name (bname ^ "_" ^ axiomsN), []),
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
   713
          [([Assumption.assume (cterm_of thy' (the asm))],
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
   714
            [(Attrib.internal o K) NewLocale.witness_attrib])])])]
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
   715
        else [];
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   716
29035
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   717
    (* These will be added in the local theory. *)
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   718
    val notes' = body_elems |>
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   719
      map (defines_to_notes thy') |>
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   720
      map (Element.morph_ctxt a_satisfy) |>
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   721
      (fn elems => fold_map assumes_to_notes elems (map Element.conclude_witness a_axioms)) |>
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   722
      fst |>
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   723
      map (Element.morph_ctxt b_satisfy) |>
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   724
      map_filter (fn Notes notes => SOME notes | _ => NONE);
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   725
b0a0843dfd99 Satisfy a_axioms.
ballarin
parents: 29034
diff changeset
   726
    val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   727
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   728
    val loc_ctxt = thy' |>
28991
694227dd3e8c dropped NameSpace.declare_base
haftmann
parents: 28965
diff changeset
   729
      NewLocale.register_locale bname (extraTs, params)
29032
3ad4cf50070d Correct order of defines in specification.
ballarin
parents: 29031
diff changeset
   730
        (asm, rev defs) ([], [])
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   731
        (map (fn n => (n, stamp ())) notes |> rev) (map (fn d => (d, stamp ())) deps' |> rev) |>
29028
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
   732
      NewLocale.init name;
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
   733
b5dad96c755a When adding locales, delay notes until local theory is built.
ballarin
parents: 29022
diff changeset
   734
  in ((name, notes'), loc_ctxt) end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   735
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   736
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   737
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   738
val add_locale_cmd = gen_add_locale read_declaration;
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   739
val add_locale = gen_add_locale cert_declaration;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   740
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   741
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   742
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   743
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   744
(*** Interpretation ***)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   745
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   746
(** Witnesses and goals **)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   747
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   748
fun prep_propp propss = propss |> map (map (rpair [] o Element.mark_witness));
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   749
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   750
fun prep_result propps thmss =
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   751
  ListPair.map (fn (props, thms) => map2 Element.make_witness props thms) (propps, thmss);
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   752
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   753
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   754
(** Interpretation between locales: declaring sublocale relationships **)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   755
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   756
local
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   757
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   758
fun gen_sublocale prep_expr intern
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   759
    raw_target expression thy =
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   760
  let
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   761
    val target = intern thy raw_target;
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   762
    val target_ctxt = NewLocale.init target thy;
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   763
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   764
    val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt;
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   765
    
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   766
    fun store_dep ((name, morph), thms) =
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   767
      NewLocale.add_dependency target (name, morph $> Element.satisfy_morphism thms $> export);
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   768
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   769
    fun after_qed results =
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   770
      ProofContext.theory (
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   771
        (* store dependencies *)
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   772
        fold store_dep (deps ~~ prep_result propss results) #>
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   773
        (* propagate registrations *)
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   774
        (fn thy => fold_rev (fn reg => NewLocale.activate_global_facts reg)
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   775
          (NewLocale.get_global_registrations thy) thy));
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   776
  in
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   777
    goal_ctxt |>
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   778
      Proof.theorem_i NONE after_qed (prep_propp propss) |>
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   779
      Element.refine_witness |> Seq.hd
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   780
  end;
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   781
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   782
in
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   783
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   784
fun sublocale_cmd x = gen_sublocale read_goal_expression NewLocale.intern x;
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   785
fun sublocale x = gen_sublocale cert_goal_expression (K I) x;
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   786
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   787
end;
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   788
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   789
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   790
(** Interpretation in theories **)
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   791
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   792
local
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   793
29211
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   794
datatype goal = Reg of string * Morphism.morphism | Eqns of Attrib.binding list;
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   795
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   796
fun gen_interpretation prep_expr parse_prop prep_attr
29210
4025459e3f83 Interpretation in theories: first version with equations.
ballarin
parents: 29208
diff changeset
   797
    expression equations thy =
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   798
  let
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   799
    val ctxt = ProofContext.init thy;
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   800
29211
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   801
    val ((propss, regs, export), expr_ctxt) = prep_expr expression ctxt;
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   802
    
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   803
    val eqns = map (parse_prop expr_ctxt o snd) equations |> Syntax.check_terms expr_ctxt;
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   804
    val eq_attns = map ((apsnd o map) (prep_attr thy) o fst) equations;
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   805
    val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt;
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   806
    val export' = Variable.export_morphism goal_ctxt expr_ctxt;
29210
4025459e3f83 Interpretation in theories: first version with equations.
ballarin
parents: 29208
diff changeset
   807
29211
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   808
    fun store (Reg (name, morph), thms) (regs, thy) =
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   809
        let
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   810
          val thms' = map (Element.morph_witness export') thms;
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   811
          val morph' = morph $> Element.satisfy_morphism thms';
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   812
          val add = NewLocale.add_global_registration (name, (morph', export));
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   813
        in ((name, morph') :: regs, add thy) end
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   814
      | store (Eqns [], []) (regs, thy) =
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   815
        let val add = fold_rev (fn (name, morph) =>
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   816
              NewLocale.activate_global_facts (name, morph $> export)) regs;
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   817
        in (regs, add thy) end
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   818
      | store (Eqns attns, thms) (regs, thy) =
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   819
        let
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   820
          val thms' = thms |> map (Element.conclude_witness #>
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   821
            Morphism.thm (export' $> export) #>
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   822
            LocalDefs.meta_rewrite_rule (ProofContext.init thy) #>
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   823
            Drule.abs_def);
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   824
          val eq_morph =
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   825
            Morphism.term_morphism (MetaSimplifier.rewrite_term thy thms' []) $>
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   826
            Morphism.thm_morphism (MetaSimplifier.rewrite_rule thms');
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   827
          val attns' = map ((apsnd o map) (Attrib.attribute_i thy)) attns;
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   828
          val add =
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   829
            fold_rev (fn (name, morph) =>
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   830
              NewLocale.amend_global_registration eq_morph (name, morph) #>
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   831
              NewLocale.activate_global_facts (name, morph $> eq_morph $> export)) regs #>
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   832
            PureThy.note_thmss Thm.lemmaK (attns' ~~ map (fn th => [([th], [])]) thms') #>
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   833
            snd
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   834
        in (regs, add thy) end;
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   835
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   836
    fun after_qed results =
29211
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   837
      ProofContext.theory (fn thy =>
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   838
        fold store (map Reg regs @ [Eqns eq_attns] ~~
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   839
          prep_result (propss @ [eqns]) results) ([], thy) |> snd);
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   840
  in
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   841
    goal_ctxt |>
29211
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   842
      Proof.theorem_i NONE after_qed (prep_propp (propss @ [eqns])) |>
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   843
      Element.refine_witness |> Seq.hd
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   844
  end;
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   845
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   846
in
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   847
29210
4025459e3f83 Interpretation in theories: first version with equations.
ballarin
parents: 29208
diff changeset
   848
fun interpretation_cmd x = gen_interpretation read_goal_expression
29211
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   849
  Syntax.parse_prop Attrib.intern_src x;
ab99da3854af Equations in interpretation as goals.
ballarin
parents: 29210
diff changeset
   850
fun interpretation x = gen_interpretation cert_goal_expression (K I) (K I) x;
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   851
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   852
end;
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   853
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   854
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   855
(** Interpretation in proof contexts **)
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   856
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   857
local
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   858
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   859
fun gen_interpret prep_expr
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   860
    expression int state =
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   861
  let
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   862
    val _ = Proof.assert_forward_or_chain state;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   863
    val ctxt = Proof.context_of state;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   864
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   865
    val ((propss, regs, export), goal_ctxt) = prep_expr expression ctxt;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   866
    
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   867
    fun store_reg ((name, morph), thms) =
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   868
      let
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   869
        val morph' = morph $> Element.satisfy_morphism thms $> export;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   870
      in
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   871
        NewLocale.activate_local_facts (name, morph')
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   872
      end;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   873
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   874
    fun after_qed results =
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   875
      Proof.map_context (fold store_reg (regs ~~ prep_result propss results)) #> Seq.single;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   876
  in
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   877
    state |> Proof.map_context (K goal_ctxt) |>
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   878
      Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   879
      "interpret" NONE after_qed (map (pair (Binding.empty, [])) (prep_propp propss)) |>
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   880
      Element.refine_witness |> Seq.hd
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   881
  end;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   882
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   883
in
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   884
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   885
fun interpret_cmd x = gen_interpret read_goal_expression x;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   886
fun interpret x = gen_interpret cert_goal_expression x;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   887
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   888
end;
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   889
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   890
end;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   891