src/Tools/isac/BridgeJEdit/Calculation.thy
author Walther Neuper <walther.neuper@jku.at>
Tue, 17 Nov 2020 14:52:57 +0100
changeset 60102 c30bba0a5f4e
parent 60101 81fbf5d64c70
child 60103 cfef3ce9bae9
permissions -rw-r--r--
testbed for SPARK parsing g_c_d.siv
walther@60096
     1
(*  Title:      src/Tools/isac/BridgeJEdit/Calculation.thy
walther@60095
     2
    Author:     Walther Neuper, JKU Linz
walther@60095
     3
    (c) due to copyright terms
walther@60095
     4
walther@60095
     5
Outer syntax for Isabelle/Isac. Compare src/Pure/Pure.thy
walther@60095
     6
*)
walther@60095
     7
walther@60095
     8
theory Calculation
walther@60101
     9
  imports "~~/src/Tools/isac/MathEngine/MathEngine" "HOL-SPARK.SPARK"
walther@60095
    10
keywords
walther@60095
    11
  "Example" :: thy_load ("str") and (* "spark_vc" :: thy_goal *)
walther@60095
    12
  "Problem" and (* root-problem + recursively in Solution *)
walther@60095
    13
  "Specification" "Model" "References" "Solution" and (* structure only *)
walther@60095
    14
  "Given" "Find" "Relate" "Where" and (* await input of term *)
walther@60095
    15
  "RTheory" (* await input of string; "R" distingues "Problem".."RProblem" *)
walther@60095
    16
  "RProblem" "RMethod" and (* await input of string list *)
walther@60095
    17
  "Tactic" (* optional for each step 0..end of calculation *)
walther@60095
    18
begin
walther@60095
    19
walther@60095
    20
ML_file parseC.sml
walther@60095
    21
walther@60099
    22
section \<open>code for spark_open: cp.to cp_spark_vcs.ML, cp_spark_vcs.ML\<close>
walther@60097
    23
text \<open>
walther@60097
    24
  We minimally change code for Example, until all works.
walther@60097
    25
walther@60097
    26
  The code is separated from Calculation.thy into files cp_spark_vcs.ML and cp_spark_commands.ML, 
walther@60097
    27
  because text\<open>...\<close> and (*...*) for const_name>\<open>...\<close> etc and @ {thms ...} causes errors.
walther@60097
    28
walther@60099
    29
  Thus for Test_Isac (**)ML_file "cp_spark_vcs.ML", ""(*cp_spark_commands.ML*), "HOL-SPARK.SPARK"
walther@60099
    30
  must be outcommented.
walther@60097
    31
\<close>
walther@60097
    32
walther@60101
    33
subsection \<open>cp. HOL/SPARK/Tools/spark_vcs.ML\<close>
walther@60097
    34
 ML \<open>
walther@60097
    35
\<close> ML \<open>
walther@60101
    36
(*  Title:      src/Tools/isac/BridgeJEdit/cp_spark_vcs.ML
walther@60101
    37
    Author:     Walther Neuper, JKU Linz
walther@60101
    38
    (c) due to copyright terms
walther@60101
    39
walther@60101
    40
Preliminary code for developing Outer_Syntax.command..Example from spark_open as a model.
walther@60101
    41
*)
walther@60101
    42
walther@60101
    43
(** theory data **)
walther@60101
    44
walther@60101
    45
fun err_unfinished () = error "An unfinished SPARK environment is still open."
walther@60101
    46
walther@60101
    47
val strip_number = apply2 implode o chop_suffix Fdl_Lexer.is_digit o raw_explode;
walther@60101
    48
walther@60101
    49
val name_ord = prod_ord string_ord (option_ord int_ord) o
walther@60101
    50
  apply2 (strip_number ##> Int.fromString);
walther@60101
    51
walther@60101
    52
structure VCtab = Table(type key = string val ord = name_ord);
walther@60101
    53
walther@60101
    54
structure VCs = Theory_Data
walther@60101
    55
(
walther@60101
    56
  type T =
walther@60101
    57
    {pfuns: ((string list * string) option * term) Symtab.table,
walther@60101
    58
     type_map: (typ * (string * string) list) Symtab.table,
walther@60101
    59
     env:
walther@60101
    60
       {ctxt: Element.context_i list,
walther@60101
    61
        defs: (binding * thm) list,
walther@60101
    62
        types: Fdl_Parser.fdl_type Fdl_Parser.tab,
walther@60101
    63
        funs: (string list * string) Fdl_Parser.tab,
walther@60101
    64
        pfuns: ((string list * string) option * term) Symtab.table,
walther@60101
    65
        ids: (term * string) Symtab.table * Name.context,
walther@60101
    66
        proving: bool,
walther@60101
    67
        vcs: (string * thm list option *
walther@60101
    68
          (string * Fdl_Parser.expr) list * (string * Fdl_Parser.expr) list) VCtab.table,
walther@60101
    69
        path: Path.T,
walther@60101
    70
        prefix: string} option}
walther@60101
    71
  val empty : T = {pfuns = Symtab.empty, type_map = Symtab.empty, env = NONE}
walther@60101
    72
  val extend = I
walther@60101
    73
  fun merge ({pfuns = pfuns1, type_map = type_map1, env = NONE},
walther@60101
    74
        {pfuns = pfuns2, type_map = type_map2, env = NONE}) =
walther@60101
    75
        {pfuns = Symtab.merge (eq_pair (op =) (op aconv)) (pfuns1, pfuns2),
walther@60101
    76
         type_map = Symtab.merge (op =) (type_map1, type_map2),
walther@60101
    77
         env = NONE}
walther@60101
    78
    | merge _ = err_unfinished ()
walther@60101
    79
)
walther@60101
    80
walther@60101
    81
(** utilities **)
walther@60101
    82
walther@60101
    83
val to_lower = raw_explode #> map Symbol.to_ascii_lower #> implode;
walther@60101
    84
walther@60101
    85
walther@60101
    86
(** set up verification conditions **)
walther@60101
    87
walther@60101
    88
fun partition_opt f =
walther@60101
    89
  let
walther@60101
    90
    fun part ys zs [] = (rev ys, rev zs)
walther@60101
    91
      | part ys zs (x :: xs) = (case f x of
walther@60101
    92
            SOME y => part (y :: ys) zs xs
walther@60101
    93
          | NONE => part ys (x :: zs) xs)
walther@60101
    94
  in part [] [] end;
walther@60101
    95
walther@60101
    96
fun dest_def (id, (Fdl_Parser.Substitution_Rule ([], Fdl_Parser.Ident s, rhs))) = SOME (id, (s, rhs))
walther@60101
    97
  | dest_def _ = NONE;
walther@60101
    98
walther@60101
    99
walther@60101
   100
val builtin = Symtab.make (map (rpair ())
walther@60101
   101
  ["->", "<->", "or", "and", "not", "=", "<>", "<", ">", "<=", ">=",
walther@60101
   102
   "+", "-", "*", "/", "div", "mod", "**"]);
walther@60101
   103
walther@60101
   104
fun complex_expr (Fdl_Parser.Number _) = false
walther@60101
   105
  | complex_expr (Fdl_Parser.Ident _) = false
walther@60101
   106
  | complex_expr (Fdl_Parser.Funct (s, es)) =
walther@60101
   107
      not (Symtab.defined builtin s) orelse exists complex_expr es
walther@60101
   108
  | complex_expr (Fdl_Parser.Quantifier (_, _, _, e)) = complex_expr e
walther@60101
   109
  | complex_expr _ = true;
walther@60101
   110
walther@60101
   111
fun complex_rule (Fdl_Parser.Inference_Rule (es, e)) =
walther@60101
   112
      complex_expr e orelse exists complex_expr es
walther@60101
   113
  | complex_rule (Fdl_Parser.Substitution_Rule (es, e, e')) =
walther@60101
   114
      complex_expr e orelse complex_expr e' orelse
walther@60101
   115
      exists complex_expr es;
walther@60101
   116
walther@60101
   117
fun is_trivial_vc ([], [(_, Fdl_Parser.Ident "true")]) = true
walther@60101
   118
  | is_trivial_vc _ = false;
walther@60101
   119
walther@60101
   120
walther@60101
   121
fun rulenames rules = commas
walther@60101
   122
  (map (fn ((s, i), _) => s ^ "(" ^ string_of_int i ^ ")") rules);
walther@60101
   123
walther@60101
   124
walther@60101
   125
val is_pfun =
walther@60101
   126
  Symtab.defined builtin orf
walther@60101
   127
  can (unprefix "fld_") orf can (unprefix "upf_") orf
walther@60101
   128
  can (unsuffix "__pos") orf can (unsuffix "__val") orf
walther@60101
   129
  equal "succ" orf equal "pred";
walther@60101
   130
walther@60101
   131
fun fold_opt f = the_default I o Option.map f;
walther@60101
   132
fun fold_pair f g (x, y) = f x #> g y;
walther@60101
   133
walther@60101
   134
fun fold_expr f g (Fdl_Parser.Funct (s, es)) = f s #> fold (fold_expr f g) es
walther@60101
   135
  | fold_expr f g (Fdl_Parser.Ident s) = g s
walther@60101
   136
  | fold_expr f g (Fdl_Parser.Number _) = I
walther@60101
   137
  | fold_expr f g (Fdl_Parser.Quantifier (_, _, _, e)) = fold_expr f g e
walther@60101
   138
  | fold_expr f g (Fdl_Parser.Element (e, es)) =
walther@60101
   139
      fold_expr f g e #> fold (fold_expr f g) es
walther@60101
   140
  | fold_expr f g (Fdl_Parser.Update (e, es, e')) =
walther@60101
   141
      fold_expr f g e #> fold (fold_expr f g) es #> fold_expr f g e'
walther@60101
   142
  | fold_expr f g (Fdl_Parser.Record (_, flds)) = fold (fold_expr f g o snd) flds
walther@60101
   143
  | fold_expr f g (Fdl_Parser.Array (_, default, assocs)) =
walther@60101
   144
      fold_opt (fold_expr f g) default #>
walther@60101
   145
      fold (fold_pair
walther@60101
   146
        (fold (fold (fold_pair
walther@60101
   147
          (fold_expr f g) (fold_opt (fold_expr f g)))))
walther@60101
   148
        (fold_expr f g)) assocs;
walther@60101
   149
walther@60101
   150
walther@60101
   151
fun add_expr_pfuns funs = fold_expr
walther@60101
   152
  (fn s => if is_pfun s then I else insert (op =) s)
walther@60101
   153
  (fn s => if is_none (Fdl_Parser.lookup funs s) then I else insert (op =) s);
walther@60101
   154
walther@60101
   155
fun lookup_prfx "" tab s = Symtab.lookup tab s
walther@60101
   156
  | lookup_prfx prfx tab s = (case Symtab.lookup tab s of
walther@60101
   157
        NONE => Symtab.lookup tab (prfx ^ "__" ^ s)
walther@60101
   158
      | x => x);
walther@60101
   159
walther@60101
   160
fun fold_vcs f vcs =
walther@60101
   161
  VCtab.fold (fn (_, (_, _, ps, cs)) => fold f ps #> fold f cs) vcs;
walther@60101
   162
walther@60101
   163
fun pfuns_of_vcs prfx funs pfuns vcs =
walther@60101
   164
  fold_vcs (add_expr_pfuns funs o snd) vcs [] |>
walther@60101
   165
  filter (is_none o lookup_prfx prfx pfuns);
walther@60101
   166
walther@60101
   167
val booleanN = "boolean";
walther@60101
   168
val integerN = "integer";
walther@60101
   169
walther@60101
   170
walther@60101
   171
fun get_type thy prfx ty =
walther@60101
   172
  let val {type_map, ...} = VCs.get thy
walther@60101
   173
  in lookup_prfx prfx type_map ty end;
walther@60101
   174
walther@60101
   175
fun mk_type' _ _ "integer" = (HOLogic.intT, [])
walther@60101
   176
  | mk_type' _ _ "boolean" = (HOLogic.boolT, [])
walther@60101
   177
  | mk_type' thy prfx ty =
walther@60101
   178
      (case get_type thy prfx ty of
walther@60101
   179
         NONE =>
walther@60101
   180
           (Syntax.check_typ (Proof_Context.init_global thy)
walther@60101
   181
              (Type (Sign.full_name thy (Binding.name ty), [])),
walther@60101
   182
            [])
walther@60101
   183
       | SOME p => p);
walther@60101
   184
walther@60101
   185
fun mk_type thy prfx ty = fst (mk_type' thy prfx ty);
walther@60101
   186
walther@60101
   187
fun check_no_assoc thy prfx s = case get_type thy prfx s of
walther@60101
   188
    NONE => ()
walther@60101
   189
  | SOME _ => error ("Cannot associate a type with " ^ s ^
walther@60101
   190
      "\nsince it is no record or enumeration type");
walther@60101
   191
walther@60101
   192
walther@60101
   193
fun define_overloaded (def_name, eq) lthy =
walther@60101
   194
  let
walther@60101
   195
    val ((c, _), rhs) = eq |> Syntax.check_term lthy |>
walther@60101
   196
      Logic.dest_equals |>> dest_Free;
walther@60101
   197
    val ((_, (_, thm)), lthy') = Local_Theory.define
walther@60101
   198
      ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs)) lthy
walther@60101
   199
    val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy');
walther@60101
   200
    val thm' = singleton (Proof_Context.export lthy' ctxt_thy) thm
walther@60101
   201
  in (thm', lthy') end;
walther@60101
   202
walther@60101
   203
(** generate properties of enumeration types **)
walther@60101
   204
walther@60101
   205
fun add_enum_type tyname tyname' thy =
walther@60101
   206
  let
walther@60101
   207
    val {case_name, ...} = the (BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] tyname');
walther@60101
   208
    val cs = map Const (the (BNF_LFP_Compat.get_constrs thy tyname'));
walther@60101
   209
    val k = length cs;
walther@60101
   210
    val T = Type (tyname', []);
walther@60101
   211
    val p = Const (\<^const_name>\<open>pos\<close>, T --> HOLogic.intT);
walther@60101
   212
    val v = Const (\<^const_name>\<open>val\<close>, HOLogic.intT --> T);
walther@60101
   213
    val card = Const (\<^const_name>\<open>card\<close>,
walther@60101
   214
      HOLogic.mk_setT T --> HOLogic.natT) $ HOLogic.mk_UNIV T;
walther@60101
   215
walther@60101
   216
    fun mk_binrel_def s f = Logic.mk_equals
walther@60101
   217
      (Const (s, T --> T --> HOLogic.boolT),
walther@60101
   218
       Abs ("x", T, Abs ("y", T,
walther@60101
   219
         Const (s, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $
walther@60101
   220
           (f $ Bound 1) $ (f $ Bound 0))));
walther@60101
   221
walther@60101
   222
    val (((def1, def2), def3), lthy) = thy |>
walther@60101
   223
walther@60101
   224
      Class.instantiation ([tyname'], [], \<^sort>\<open>spark_enum\<close>) |>
walther@60101
   225
walther@60101
   226
      define_overloaded ("pos_" ^ tyname ^ "_def", Logic.mk_equals
walther@60101
   227
        (p,
walther@60101
   228
         list_comb (Const (case_name, replicate k HOLogic.intT @
walther@60101
   229
             [T] ---> HOLogic.intT),
walther@60101
   230
           map (HOLogic.mk_number HOLogic.intT) (0 upto k - 1)))) ||>>
walther@60101
   231
walther@60101
   232
      define_overloaded ("less_eq_" ^ tyname ^ "_def",
walther@60101
   233
        mk_binrel_def \<^const_name>\<open>less_eq\<close> p) ||>>
walther@60101
   234
      define_overloaded ("less_" ^ tyname ^ "_def",
walther@60101
   235
        mk_binrel_def \<^const_name>\<open>less\<close> p);
walther@60101
   236
walther@60101
   237
    val UNIV_eq = Goal.prove lthy [] []
walther@60101
   238
      (HOLogic.mk_Trueprop (HOLogic.mk_eq
walther@60101
   239
         (HOLogic.mk_UNIV T, HOLogic.mk_set T cs)))
walther@60101
   240
      (fn {context = ctxt, ...} =>
walther@60101
   241
         resolve_tac ctxt @{thms subset_antisym} 1 THEN
walther@60101
   242
         resolve_tac ctxt @{thms subsetI} 1 THEN
walther@60101
   243
         Old_Datatype_Aux.exh_tac ctxt (K (#exhaust (BNF_LFP_Compat.the_info
walther@60101
   244
           (Proof_Context.theory_of ctxt) [BNF_LFP_Compat.Keep_Nesting] tyname'))) 1 THEN
walther@60101
   245
         ALLGOALS (asm_full_simp_tac ctxt));
walther@60101
   246
walther@60101
   247
    val finite_UNIV = Goal.prove lthy [] []
walther@60101
   248
      (HOLogic.mk_Trueprop (Const (\<^const_name>\<open>finite\<close>,
walther@60101
   249
         HOLogic.mk_setT T --> HOLogic.boolT) $ HOLogic.mk_UNIV T))
walther@60101
   250
      (fn {context = ctxt, ...} => simp_tac (ctxt addsimps [UNIV_eq]) 1);
walther@60101
   251
walther@60101
   252
    val card_UNIV = Goal.prove lthy [] []
walther@60101
   253
      (HOLogic.mk_Trueprop (HOLogic.mk_eq
walther@60101
   254
         (card, HOLogic.mk_number HOLogic.natT k)))
walther@60101
   255
      (fn {context = ctxt, ...} => simp_tac (ctxt addsimps [UNIV_eq]) 1);
walther@60101
   256
walther@60101
   257
    val range_pos = Goal.prove lthy [] []
walther@60101
   258
      (HOLogic.mk_Trueprop (HOLogic.mk_eq
walther@60101
   259
         (Const (\<^const_name>\<open>image\<close>, (T --> HOLogic.intT) -->
walther@60101
   260
            HOLogic.mk_setT T --> HOLogic.mk_setT HOLogic.intT) $
walther@60101
   261
              p $ HOLogic.mk_UNIV T,
walther@60101
   262
          Const (\<^const_name>\<open>atLeastLessThan\<close>, HOLogic.intT -->
walther@60101
   263
            HOLogic.intT --> HOLogic.mk_setT HOLogic.intT) $
walther@60101
   264
              HOLogic.mk_number HOLogic.intT 0 $
walther@60101
   265
              (\<^term>\<open>int\<close> $ card))))
walther@60101
   266
      (fn {context = ctxt, ...} =>
walther@60101
   267
         simp_tac (ctxt addsimps [card_UNIV]) 1 THEN
walther@60101
   268
         simp_tac (ctxt addsimps [UNIV_eq, def1]) 1 THEN
walther@60101
   269
         resolve_tac ctxt @{thms subset_antisym} 1 THEN
walther@60101
   270
         simp_tac ctxt 1 THEN
walther@60101
   271
         resolve_tac ctxt @{thms subsetI} 1 THEN
walther@60101
   272
         asm_full_simp_tac (ctxt addsimps @{thms interval_expand}
walther@60101
   273
           delsimps @{thms atLeastLessThan_iff}) 1);
walther@60101
   274
walther@60101
   275
    val lthy' =
walther@60101
   276
      Class.prove_instantiation_instance (fn ctxt =>
walther@60101
   277
        Class.intro_classes_tac ctxt [] THEN
walther@60101
   278
        resolve_tac ctxt [finite_UNIV] 1 THEN
walther@60101
   279
        resolve_tac ctxt [range_pos] 1 THEN
walther@60101
   280
        simp_tac (put_simpset HOL_basic_ss ctxt addsimps [def3]) 1 THEN
walther@60101
   281
        simp_tac (put_simpset HOL_basic_ss ctxt addsimps [def2]) 1) lthy;
walther@60101
   282
walther@60101
   283
    val (pos_eqs, val_eqs) = split_list (map_index (fn (i, c) =>
walther@60101
   284
      let
walther@60101
   285
        val n = HOLogic.mk_number HOLogic.intT i;
walther@60101
   286
        val th = Goal.prove lthy' [] []
walther@60101
   287
          (HOLogic.mk_Trueprop (HOLogic.mk_eq (p $ c, n)))
walther@60101
   288
          (fn {context = ctxt, ...} => simp_tac (ctxt addsimps [def1]) 1);
walther@60101
   289
        val th' = Goal.prove lthy' [] []
walther@60101
   290
          (HOLogic.mk_Trueprop (HOLogic.mk_eq (v $ n, c)))
walther@60101
   291
          (fn {context = ctxt, ...} =>
walther@60101
   292
             resolve_tac ctxt [@{thm inj_pos} RS @{thm injD}] 1 THEN
walther@60101
   293
             simp_tac (ctxt addsimps [@{thm pos_val}, range_pos, card_UNIV, th]) 1)
walther@60101
   294
      in (th, th') end) cs);
walther@60101
   295
walther@60101
   296
    val first_el = Goal.prove lthy' [] []
walther@60101
   297
      (HOLogic.mk_Trueprop (HOLogic.mk_eq
walther@60101
   298
         (Const (\<^const_name>\<open>first_el\<close>, T), hd cs)))
walther@60101
   299
      (fn {context = ctxt, ...} => simp_tac (ctxt addsimps [@{thm first_el_def}, hd val_eqs]) 1);
walther@60101
   300
walther@60101
   301
    val last_el = Goal.prove lthy' [] []
walther@60101
   302
      (HOLogic.mk_Trueprop (HOLogic.mk_eq
walther@60101
   303
         (Const (\<^const_name>\<open>last_el\<close>, T), List.last cs)))
walther@60101
   304
      (fn {context = ctxt, ...} =>
walther@60101
   305
        simp_tac (ctxt addsimps [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1);
walther@60101
   306
  in
walther@60101
   307
    lthy' |>
walther@60101
   308
    Local_Theory.note
walther@60101
   309
      ((Binding.name (tyname ^ "_card"), @{attributes [simp]}), [card_UNIV]) ||>>
walther@60101
   310
    Local_Theory.note
walther@60101
   311
      ((Binding.name (tyname ^ "_pos"), @{attributes [simp]}), pos_eqs) ||>>
walther@60101
   312
    Local_Theory.note
walther@60101
   313
      ((Binding.name (tyname ^ "_val"), @{attributes [simp]}), val_eqs) ||>>
walther@60101
   314
    Local_Theory.note
walther@60101
   315
      ((Binding.name (tyname ^ "_first_el"), @{attributes [simp]}), [first_el]) ||>>
walther@60101
   316
    Local_Theory.note
walther@60101
   317
      ((Binding.name (tyname ^ "_last_el"), @{attributes [simp]}), [last_el]) |> snd |>
walther@60101
   318
    Local_Theory.exit_global
walther@60101
   319
  end;
walther@60101
   320
walther@60101
   321
walther@60101
   322
val lcase_eq = (op =) o apply2 (to_lower o Long_Name.base_name);
walther@60101
   323
walther@60101
   324
fun check_enum [] [] = NONE
walther@60101
   325
  | check_enum els [] = SOME ("has no element(s) " ^ commas els)
walther@60101
   326
  | check_enum [] cs = SOME ("has extra element(s) " ^
walther@60101
   327
      commas (map (Long_Name.base_name o fst) cs))
walther@60101
   328
  | check_enum (el :: els) ((cname, _) :: cs) =
walther@60101
   329
      if lcase_eq (el, cname) then check_enum els cs
walther@60101
   330
      else SOME ("either has no element " ^ el ^
walther@60101
   331
        " or it is at the wrong position");
walther@60101
   332
walther@60101
   333
fun strip_prfx s =
walther@60101
   334
  let
walther@60101
   335
    fun strip ys [] = ("", implode ys)
walther@60101
   336
      | strip ys ("_" :: "_" :: xs) = (implode (rev xs), implode ys)
walther@60101
   337
      | strip ys (x :: xs) = strip (x :: ys) xs
walther@60101
   338
  in strip [] (rev (raw_explode s)) end;
walther@60101
   339
walther@60101
   340
fun assoc_ty_err thy T s msg =
walther@60101
   341
  error ("Type " ^ Syntax.string_of_typ_global thy T ^
walther@60101
   342
    " associated with SPARK type " ^ s ^ "\n" ^ msg);
walther@60101
   343
walther@60101
   344
fun check_enum [] [] = NONE
walther@60101
   345
  | check_enum els [] = SOME ("has no element(s) " ^ commas els)
walther@60101
   346
  | check_enum [] cs = SOME ("has extra element(s) " ^
walther@60101
   347
      commas (map (Long_Name.base_name o fst) cs))
walther@60101
   348
  | check_enum (el :: els) ((cname, _) :: cs) =
walther@60101
   349
      if lcase_eq (el, cname) then check_enum els cs
walther@60101
   350
      else SOME ("either has no element " ^ el ^
walther@60101
   351
        " or it is at the wrong position");
walther@60101
   352
walther@60101
   353
fun unprefix_pfun "" s = s
walther@60101
   354
  | unprefix_pfun prfx s =
walther@60101
   355
      let val (prfx', s') = strip_prfx s
walther@60101
   356
      in if prfx = prfx' then s' else s end;
walther@60101
   357
walther@60101
   358
fun invert_map [] = I
walther@60101
   359
  | invert_map cmap =
walther@60101
   360
      map (apfst (the o AList.lookup (op =) (map swap cmap)));
walther@60101
   361
walther@60101
   362
fun get_record_info thy T = (case Record.dest_recTs T of
walther@60101
   363
    [(tyname, [\<^typ>\<open>unit\<close>])] =>
walther@60101
   364
      Record.get_info thy (Long_Name.qualifier tyname)
walther@60101
   365
  | _ => NONE);
walther@60101
   366
walther@60101
   367
walther@60101
   368
walther@60101
   369
fun add_type_def prfx (s, Fdl_Parser.Basic_Type ty) (ids, thy) =
walther@60101
   370
      (check_no_assoc thy prfx s;
walther@60101
   371
       (ids,
walther@60101
   372
        Typedecl.abbrev_global (Binding.name s, [], NoSyn)
walther@60101
   373
          (mk_type thy prfx ty) thy |> snd))
walther@60101
   374
walther@60101
   375
  | add_type_def prfx (s, Fdl_Parser.Enum_Type els) ((tab, ctxt), thy) =
walther@60101
   376
      let
walther@60101
   377
        val (thy', tyname) = (case get_type thy prfx s of
walther@60101
   378
            NONE =>
walther@60101
   379
              let
walther@60101
   380
                val tyb = Binding.name s;
walther@60101
   381
                val tyname = Sign.full_name thy tyb
walther@60101
   382
              in
walther@60101
   383
                (thy |>
walther@60101
   384
                 BNF_LFP_Compat.add_datatype [BNF_LFP_Compat.Keep_Nesting]
walther@60101
   385
                   [((tyb, [], NoSyn),
walther@60101
   386
                     map (fn s => (Binding.name s, [], NoSyn)) els)] |> snd |>
walther@60101
   387
                 add_enum_type s tyname,
walther@60101
   388
                 tyname)
walther@60101
   389
              end
walther@60101
   390
          | SOME (T as Type (tyname, []), cmap) =>
walther@60101
   391
              (case BNF_LFP_Compat.get_constrs thy tyname of
walther@60101
   392
                 NONE => assoc_ty_err thy T s "is not a datatype"
walther@60101
   393
               | SOME cs =>
walther@60101
   394
                   let val (prfx', _) = strip_prfx s
walther@60101
   395
                   in
walther@60101
   396
                     case check_enum (map (unprefix_pfun prfx') els)
walther@60101
   397
                         (invert_map cmap cs) of
walther@60101
   398
                       NONE => (thy, tyname)
walther@60101
   399
                     | SOME msg => assoc_ty_err thy T s msg
walther@60101
   400
                   end)
walther@60101
   401
          | SOME (T, _) => assoc_ty_err thy T s "is not a datatype");
walther@60101
   402
        val cs = map Const (the (BNF_LFP_Compat.get_constrs thy' tyname));
walther@60101
   403
      in
walther@60101
   404
        ((fold (Symtab.update_new o apsnd (rpair s)) (els ~~ cs) tab,
walther@60101
   405
          fold Name.declare els ctxt),
walther@60101
   406
         thy')
walther@60101
   407
      end
walther@60101
   408
walther@60101
   409
  | add_type_def prfx (s, Fdl_Parser.Array_Type (argtys, resty)) (ids, thy) =
walther@60101
   410
      (check_no_assoc thy prfx s;
walther@60101
   411
       (ids,
walther@60101
   412
        Typedecl.abbrev_global (Binding.name s, [], NoSyn)
walther@60101
   413
          (foldr1 HOLogic.mk_prodT (map (mk_type thy prfx) argtys) -->
walther@60101
   414
             mk_type thy prfx resty) thy |> snd))
walther@60101
   415
walther@60101
   416
  | add_type_def prfx (s, Fdl_Parser.Record_Type fldtys) (ids, thy) =
walther@60101
   417
      (ids,
walther@60101
   418
       let val fldTs = maps (fn (flds, ty) =>
walther@60101
   419
         map (rpair (mk_type thy prfx ty)) flds) fldtys
walther@60101
   420
       in case get_type thy prfx s of
walther@60101
   421
           NONE =>
walther@60101
   422
             Record.add_record {overloaded = false} ([], Binding.name s) NONE
walther@60101
   423
               (map (fn (fld, T) => (Binding.name fld, T, NoSyn)) fldTs) thy
walther@60101
   424
         | SOME (rT, cmap) =>
walther@60101
   425
             (case get_record_info thy rT of
walther@60101
   426
                NONE => assoc_ty_err thy rT s "is not a record type"
walther@60101
   427
              | SOME {fields, ...} =>
walther@60101
   428
                  let val fields' = invert_map cmap fields
walther@60101
   429
                  in
walther@60101
   430
                    (case subtract (lcase_eq o apply2 fst) fldTs fields' of
walther@60101
   431
                       [] => ()
walther@60101
   432
                     | flds => assoc_ty_err thy rT s ("has extra field(s) " ^
walther@60101
   433
                         commas (map (Long_Name.base_name o fst) flds));
walther@60101
   434
                     map (fn (fld, T) =>
walther@60101
   435
                       case AList.lookup lcase_eq fields' fld of
walther@60101
   436
                         NONE => assoc_ty_err thy rT s ("has no field " ^ fld)
walther@60101
   437
                       | SOME U => T = U orelse assoc_ty_err thy rT s
walther@60101
   438
                           ("has field " ^
walther@60101
   439
                            fld ^ " whose type\n" ^
walther@60101
   440
                            Syntax.string_of_typ_global thy U ^
walther@60101
   441
                            "\ndoes not match declared type\n" ^
walther@60101
   442
                            Syntax.string_of_typ_global thy T)) fldTs;
walther@60101
   443
                     thy)
walther@60101
   444
                  end)
walther@60101
   445
       end)
walther@60101
   446
walther@60101
   447
  | add_type_def prfx (s, Fdl_Parser.Pending_Type) (ids, thy) =
walther@60101
   448
      (ids,
walther@60101
   449
       case get_type thy prfx s of
walther@60101
   450
         SOME _ => thy
walther@60101
   451
       | NONE => Typedecl.typedecl_global {final = true} (Binding.name s, [], NoSyn) thy |> snd);
walther@60101
   452
walther@60101
   453
fun add_type_def prfx (s, Fdl_Parser.Basic_Type ty) (ids, thy) =
walther@60101
   454
      (check_no_assoc thy prfx s;
walther@60101
   455
       (ids,
walther@60101
   456
        Typedecl.abbrev_global (Binding.name s, [], NoSyn)
walther@60101
   457
          (mk_type thy prfx ty) thy |> snd))
walther@60101
   458
walther@60101
   459
  | add_type_def prfx (s, Fdl_Parser.Enum_Type els) ((tab, ctxt), thy) =
walther@60101
   460
      let
walther@60101
   461
        val (thy', tyname) = (case get_type thy prfx s of
walther@60101
   462
            NONE =>
walther@60101
   463
              let
walther@60101
   464
                val tyb = Binding.name s;
walther@60101
   465
                val tyname = Sign.full_name thy tyb
walther@60101
   466
              in
walther@60101
   467
                (thy |>
walther@60101
   468
                 BNF_LFP_Compat.add_datatype [BNF_LFP_Compat.Keep_Nesting]
walther@60101
   469
                   [((tyb, [], NoSyn),
walther@60101
   470
                     map (fn s => (Binding.name s, [], NoSyn)) els)] |> snd |>
walther@60101
   471
                 add_enum_type s tyname,
walther@60101
   472
                 tyname)
walther@60101
   473
              end
walther@60101
   474
          | SOME (T as Type (tyname, []), cmap) =>
walther@60101
   475
              (case BNF_LFP_Compat.get_constrs thy tyname of
walther@60101
   476
                 NONE => assoc_ty_err thy T s "is not a datatype"
walther@60101
   477
               | SOME cs =>
walther@60101
   478
                   let val (prfx', _) = strip_prfx s
walther@60101
   479
                   in
walther@60101
   480
                     case check_enum (map (unprefix_pfun prfx') els)
walther@60101
   481
                         (invert_map cmap cs) of
walther@60101
   482
                       NONE => (thy, tyname)
walther@60101
   483
                     | SOME msg => assoc_ty_err thy T s msg
walther@60101
   484
                   end)
walther@60101
   485
          | SOME (T, _) => assoc_ty_err thy T s "is not a datatype");
walther@60101
   486
        val cs = map Const (the (BNF_LFP_Compat.get_constrs thy' tyname));
walther@60101
   487
      in
walther@60101
   488
        ((fold (Symtab.update_new o apsnd (rpair s)) (els ~~ cs) tab,
walther@60101
   489
          fold Name.declare els ctxt),
walther@60101
   490
         thy')
walther@60101
   491
      end
walther@60101
   492
walther@60101
   493
  | add_type_def prfx (s, Fdl_Parser.Array_Type (argtys, resty)) (ids, thy) =
walther@60101
   494
      (check_no_assoc thy prfx s;
walther@60101
   495
       (ids,
walther@60101
   496
        Typedecl.abbrev_global (Binding.name s, [], NoSyn)
walther@60101
   497
          (foldr1 HOLogic.mk_prodT (map (mk_type thy prfx) argtys) -->
walther@60101
   498
             mk_type thy prfx resty) thy |> snd))
walther@60101
   499
walther@60101
   500
  | add_type_def prfx (s, Fdl_Parser.Record_Type fldtys) (ids, thy) =
walther@60101
   501
      (ids,
walther@60101
   502
       let val fldTs = maps (fn (flds, ty) =>
walther@60101
   503
         map (rpair (mk_type thy prfx ty)) flds) fldtys
walther@60101
   504
       in case get_type thy prfx s of
walther@60101
   505
           NONE =>
walther@60101
   506
             Record.add_record {overloaded = false} ([], Binding.name s) NONE
walther@60101
   507
               (map (fn (fld, T) => (Binding.name fld, T, NoSyn)) fldTs) thy
walther@60101
   508
         | SOME (rT, cmap) =>
walther@60101
   509
             (case get_record_info thy rT of
walther@60101
   510
                NONE => assoc_ty_err thy rT s "is not a record type"
walther@60101
   511
              | SOME {fields, ...} =>
walther@60101
   512
                  let val fields' = invert_map cmap fields
walther@60101
   513
                  in
walther@60101
   514
                    (case subtract (lcase_eq o apply2 fst) fldTs fields' of
walther@60101
   515
                       [] => ()
walther@60101
   516
                     | flds => assoc_ty_err thy rT s ("has extra field(s) " ^
walther@60101
   517
                         commas (map (Long_Name.base_name o fst) flds));
walther@60101
   518
                     map (fn (fld, T) =>
walther@60101
   519
                       case AList.lookup lcase_eq fields' fld of
walther@60101
   520
                         NONE => assoc_ty_err thy rT s ("has no field " ^ fld)
walther@60101
   521
                       | SOME U => T = U orelse assoc_ty_err thy rT s
walther@60101
   522
                           ("has field " ^
walther@60101
   523
                            fld ^ " whose type\n" ^
walther@60101
   524
                            Syntax.string_of_typ_global thy U ^
walther@60101
   525
                            "\ndoes not match declared type\n" ^
walther@60101
   526
                            Syntax.string_of_typ_global thy T)) fldTs;
walther@60101
   527
                     thy)
walther@60101
   528
                  end)
walther@60101
   529
       end)
walther@60101
   530
walther@60101
   531
  | add_type_def prfx (s, Fdl_Parser.Pending_Type) (ids, thy) =
walther@60101
   532
      (ids,
walther@60101
   533
       case get_type thy prfx s of
walther@60101
   534
         SOME _ => thy
walther@60101
   535
       | NONE => Typedecl.typedecl_global {final = true} (Binding.name s, [], NoSyn) thy |> snd);
walther@60101
   536
walther@60101
   537
fun add_const prfx (s, ty) ((tab, ctxt), thy) =
walther@60101
   538
  let
walther@60101
   539
    val T = mk_type thy prfx ty;
walther@60101
   540
    val b = Binding.name s;
walther@60101
   541
    val c = Const (Sign.full_name thy b, T)
walther@60101
   542
  in
walther@60101
   543
    (c,
walther@60101
   544
     ((Symtab.update (s, (c, ty)) tab, Name.declare s ctxt),
walther@60101
   545
      Sign.add_consts [(b, T, NoSyn)] thy))
walther@60101
   546
  end;
walther@60101
   547
walther@60101
   548
fun mk_unop s t =
walther@60101
   549
  let val T = fastype_of t
walther@60101
   550
  in Const (s, T --> T) $ t end;
walther@60101
   551
walther@60101
   552
fun strip_underscores s =
walther@60101
   553
  strip_underscores (unsuffix "_" s) handle Fail _ => s;
walther@60101
   554
walther@60101
   555
fun strip_tilde s =
walther@60101
   556
  unsuffix "~" s ^ "_init" handle Fail _ => s;
walther@60101
   557
walther@60101
   558
val mangle_name = strip_underscores #> strip_tilde;
walther@60101
   559
walther@60101
   560
fun mk_variables thy prfx xs ty (tab, ctxt) =
walther@60101
   561
  let
walther@60101
   562
    val T = mk_type thy prfx ty;
walther@60101
   563
    val (ys, ctxt') = fold_map Name.variant (map mangle_name xs) ctxt;
walther@60101
   564
    val zs = map (Free o rpair T) ys;
walther@60101
   565
  in (zs, (fold (Symtab.update o apsnd (rpair ty)) (xs ~~ zs) tab, ctxt')) end;
walther@60101
   566
walther@60101
   567
fun find_field [] fname fields =
walther@60101
   568
      find_first (curry lcase_eq fname o fst) fields
walther@60101
   569
  | find_field cmap fname fields =
walther@60101
   570
      (case AList.lookup (op =) cmap fname of
walther@60101
   571
         NONE => NONE
walther@60101
   572
       | SOME fname' => SOME (fname', the (AList.lookup (op =) fields fname')));
walther@60101
   573
walther@60101
   574
fun find_field' fname = get_first (fn (flds, fldty) =>
walther@60101
   575
  if member (op =) flds fname then SOME fldty else NONE);
walther@60101
   576
walther@60101
   577
fun mk_times (t, u) =
walther@60101
   578
  let
walther@60101
   579
    val setT = fastype_of t;
walther@60101
   580
    val T = HOLogic.dest_setT setT;
walther@60101
   581
    val U = HOLogic.dest_setT (fastype_of u)
walther@60101
   582
  in
walther@60101
   583
    Const (\<^const_name>\<open>Sigma\<close>, setT --> (T --> HOLogic.mk_setT U) -->
walther@60101
   584
      HOLogic.mk_setT (HOLogic.mk_prodT (T, U))) $ t $ Abs ("", T, u)
walther@60101
   585
  end;
walther@60101
   586
walther@60101
   587
walther@60101
   588
fun term_of_expr thy prfx types pfuns =
walther@60101
   589
  let
walther@60101
   590
    fun tm_of vs (Fdl_Parser.Funct ("->", [e, e'])) =
walther@60101
   591
          (HOLogic.mk_imp (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
walther@60101
   592
walther@60101
   593
      | tm_of vs (Fdl_Parser.Funct ("<->", [e, e'])) =
walther@60101
   594
          (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
walther@60101
   595
walther@60101
   596
      | tm_of vs (Fdl_Parser.Funct ("or", [e, e'])) =
walther@60101
   597
          (HOLogic.mk_disj (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
walther@60101
   598
walther@60101
   599
      | tm_of vs (Fdl_Parser.Funct ("and", [e, e'])) =
walther@60101
   600
          (HOLogic.mk_conj (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
walther@60101
   601
walther@60101
   602
      | tm_of vs (Fdl_Parser.Funct ("not", [e])) =
walther@60101
   603
          (HOLogic.mk_not (fst (tm_of vs e)), booleanN)
walther@60101
   604
walther@60101
   605
      | tm_of vs (Fdl_Parser.Funct ("=", [e, e'])) =
walther@60101
   606
          (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
walther@60101
   607
walther@60101
   608
      | tm_of vs (Fdl_Parser.Funct ("<>", [e, e'])) = (HOLogic.mk_not
walther@60101
   609
          (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e'))), booleanN)
walther@60101
   610
walther@60101
   611
      | tm_of vs (Fdl_Parser.Funct ("<", [e, e'])) = (HOLogic.mk_binrel \<^const_name>\<open>less\<close>
walther@60101
   612
          (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
walther@60101
   613
walther@60101
   614
      | tm_of vs (Fdl_Parser.Funct (">", [e, e'])) = (HOLogic.mk_binrel \<^const_name>\<open>less\<close>
walther@60101
   615
          (fst (tm_of vs e'), fst (tm_of vs e)), booleanN)
walther@60101
   616
walther@60101
   617
      | tm_of vs (Fdl_Parser.Funct ("<=", [e, e'])) = (HOLogic.mk_binrel \<^const_name>\<open>less_eq\<close>
walther@60101
   618
          (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
walther@60101
   619
walther@60101
   620
      | tm_of vs (Fdl_Parser.Funct (">=", [e, e'])) = (HOLogic.mk_binrel \<^const_name>\<open>less_eq\<close>
walther@60101
   621
          (fst (tm_of vs e'), fst (tm_of vs e)), booleanN)
walther@60101
   622
walther@60101
   623
      | tm_of vs (Fdl_Parser.Funct ("+", [e, e'])) = (HOLogic.mk_binop \<^const_name>\<open>plus\<close>
walther@60101
   624
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
walther@60101
   625
walther@60101
   626
      | tm_of vs (Fdl_Parser.Funct ("-", [e, e'])) = (HOLogic.mk_binop \<^const_name>\<open>minus\<close>
walther@60101
   627
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
walther@60101
   628
walther@60101
   629
      | tm_of vs (Fdl_Parser.Funct ("*", [e, e'])) = (HOLogic.mk_binop \<^const_name>\<open>times\<close>
walther@60101
   630
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
walther@60101
   631
walther@60101
   632
      | tm_of vs (Fdl_Parser.Funct ("/", [e, e'])) = (HOLogic.mk_binop \<^const_name>\<open>divide\<close>
walther@60101
   633
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
walther@60101
   634
walther@60101
   635
      | tm_of vs (Fdl_Parser.Funct ("div", [e, e'])) = (HOLogic.mk_binop \<^const_name>\<open>sdiv\<close>
walther@60101
   636
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
walther@60101
   637
walther@60101
   638
      | tm_of vs (Fdl_Parser.Funct ("mod", [e, e'])) = (HOLogic.mk_binop \<^const_name>\<open>modulo\<close>
walther@60101
   639
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
walther@60101
   640
walther@60101
   641
      | tm_of vs (Fdl_Parser.Funct ("-", [e])) =
walther@60101
   642
          (mk_unop \<^const_name>\<open>uminus\<close> (fst (tm_of vs e)), integerN)
walther@60101
   643
walther@60101
   644
      | tm_of vs (Fdl_Parser.Funct ("**", [e, e'])) =
walther@60101
   645
          (Const (\<^const_name>\<open>power\<close>, HOLogic.intT --> HOLogic.natT -->
walther@60101
   646
             HOLogic.intT) $ fst (tm_of vs e) $
walther@60101
   647
               (\<^const>\<open>nat\<close> $ fst (tm_of vs e')), integerN)
walther@60101
   648
walther@60101
   649
      | tm_of (tab, _) (Fdl_Parser.Ident s) =
walther@60101
   650
          (case Symtab.lookup tab s of
walther@60101
   651
             SOME t_ty => t_ty
walther@60101
   652
           | NONE => (case lookup_prfx prfx pfuns s of
walther@60101
   653
               SOME (SOME ([], resty), t) => (t, resty)
walther@60101
   654
             | _ => error ("Undeclared identifier " ^ s)))
walther@60101
   655
walther@60101
   656
      | tm_of _ (Fdl_Parser.Number i) = (HOLogic.mk_number HOLogic.intT i, integerN)
walther@60101
   657
walther@60101
   658
      | tm_of vs (Fdl_Parser.Quantifier (s, xs, ty, e)) =
walther@60101
   659
          let
walther@60101
   660
            val (ys, vs') = mk_variables thy prfx xs ty vs;
walther@60101
   661
            val q = (case s of
walther@60101
   662
                "for_all" => HOLogic.mk_all
walther@60101
   663
              | "for_some" => HOLogic.mk_exists)
walther@60101
   664
          in
walther@60101
   665
            (fold_rev (fn Free (x, T) => fn t => q (x, T, t))
walther@60101
   666
               ys (fst (tm_of vs' e)),
walther@60101
   667
             booleanN)
walther@60101
   668
          end
walther@60101
   669
walther@60101
   670
      | tm_of vs (Fdl_Parser.Funct (s, es)) =
walther@60101
   671
walther@60101
   672
          (* record field selection *)
walther@60101
   673
          (case try (unprefix "fld_") s of
walther@60101
   674
             SOME fname => (case es of
walther@60101
   675
               [e] =>
walther@60101
   676
                 let
walther@60101
   677
                   val (t, rcdty) = tm_of vs e;
walther@60101
   678
                   val (rT, cmap) = mk_type' thy prfx rcdty
walther@60101
   679
                 in case (get_record_info thy rT, Fdl_Parser.lookup types rcdty) of
walther@60101
   680
                     (SOME {fields, ...}, SOME (Fdl_Parser.Record_Type fldtys)) =>
walther@60101
   681
                       (case (find_field cmap fname fields,
walther@60101
   682
                            find_field' fname fldtys) of
walther@60101
   683
                          (SOME (fname', fT), SOME fldty) =>
walther@60101
   684
                            (Const (fname', rT --> fT) $ t, fldty)
walther@60101
   685
                        | _ => error ("Record " ^ rcdty ^
walther@60101
   686
                            " has no field named " ^ fname))
walther@60101
   687
                   | _ => error (rcdty ^ " is not a record type")
walther@60101
   688
                 end
walther@60101
   689
             | _ => error ("Function " ^ s ^ " expects one argument"))
walther@60101
   690
           | NONE =>
walther@60101
   691
walther@60101
   692
          (* record field update *)
walther@60101
   693
          (case try (unprefix "upf_") s of
walther@60101
   694
             SOME fname => (case es of
walther@60101
   695
               [e, e'] =>
walther@60101
   696
                 let
walther@60101
   697
                   val (t, rcdty) = tm_of vs e;
walther@60101
   698
                   val (rT, cmap) = mk_type' thy prfx rcdty;
walther@60101
   699
                   val (u, fldty) = tm_of vs e';
walther@60101
   700
                   val fT = mk_type thy prfx fldty
walther@60101
   701
                 in case get_record_info thy rT of
walther@60101
   702
                     SOME {fields, ...} =>
walther@60101
   703
                       (case find_field cmap fname fields of
walther@60101
   704
                          SOME (fname', fU) =>
walther@60101
   705
                            if fT = fU then
walther@60101
   706
                              (Const (fname' ^ "_update",
walther@60101
   707
                                 (fT --> fT) --> rT --> rT) $
walther@60101
   708
                                   Abs ("x", fT, u) $ t,
walther@60101
   709
                               rcdty)
walther@60101
   710
                            else error ("Type\n" ^
walther@60101
   711
                              Syntax.string_of_typ_global thy fT ^
walther@60101
   712
                              "\ndoes not match type\n" ^
walther@60101
   713
                              Syntax.string_of_typ_global thy fU ^
walther@60101
   714
                              "\nof field " ^ fname)
walther@60101
   715
                        | NONE => error ("Record " ^ rcdty ^
walther@60101
   716
                            " has no field named " ^ fname))
walther@60101
   717
                   | _ => error (rcdty ^ " is not a record type")
walther@60101
   718
                 end
walther@60101
   719
             | _ => error ("Function " ^ s ^ " expects two arguments"))
walther@60101
   720
           | NONE =>
walther@60101
   721
walther@60101
   722
          (* enumeration type to integer *)
walther@60101
   723
          (case try (unsuffix "__pos") s of
walther@60101
   724
             SOME tyname => (case es of
walther@60101
   725
               [e] => (Const (\<^const_name>\<open>pos\<close>,
walther@60101
   726
                   mk_type thy prfx tyname --> HOLogic.intT) $ fst (tm_of vs e),
walther@60101
   727
                 integerN)
walther@60101
   728
             | _ => error ("Function " ^ s ^ " expects one argument"))
walther@60101
   729
           | NONE =>
walther@60101
   730
walther@60101
   731
          (* integer to enumeration type *)
walther@60101
   732
          (case try (unsuffix "__val") s of
walther@60101
   733
             SOME tyname => (case es of
walther@60101
   734
               [e] => (Const (\<^const_name>\<open>val\<close>,
walther@60101
   735
                   HOLogic.intT --> mk_type thy prfx tyname) $ fst (tm_of vs e),
walther@60101
   736
                 tyname)
walther@60101
   737
             | _ => error ("Function " ^ s ^ " expects one argument"))
walther@60101
   738
           | NONE =>
walther@60101
   739
walther@60101
   740
          (* successor / predecessor of enumeration type element *)
walther@60101
   741
          if s = "succ" orelse s = "pred" then (case es of
walther@60101
   742
              [e] =>
walther@60101
   743
                let
walther@60101
   744
                  val (t, tyname) = tm_of vs e;
walther@60101
   745
                  val T = mk_type thy prfx tyname
walther@60101
   746
                in (Const
walther@60101
   747
                  (if s = "succ" then \<^const_name>\<open>succ\<close>
walther@60101
   748
                   else \<^const_name>\<open>pred\<close>, T --> T) $ t, tyname)
walther@60101
   749
                end
walther@60101
   750
            | _ => error ("Function " ^ s ^ " expects one argument"))
walther@60101
   751
walther@60101
   752
          (* user-defined proof function *)
walther@60101
   753
          else
walther@60101
   754
            (case lookup_prfx prfx pfuns s of
walther@60101
   755
               SOME (SOME (_, resty), t) =>
walther@60101
   756
                 (list_comb (t, map (fst o tm_of vs) es), resty)
walther@60101
   757
             | _ => error ("Undeclared proof function " ^ s))))))
walther@60101
   758
walther@60101
   759
      | tm_of vs (Fdl_Parser.Element (e, es)) =
walther@60101
   760
          let val (t, ty) = tm_of vs e
walther@60101
   761
          in case Fdl_Parser.lookup types ty of
walther@60101
   762
              SOME (Fdl_Parser.Array_Type (_, elty)) =>
walther@60101
   763
                (t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es), elty)
walther@60101
   764
            | _ => error (ty ^ " is not an array type")
walther@60101
   765
          end
walther@60101
   766
walther@60101
   767
      | tm_of vs (Fdl_Parser.Update (e, es, e')) =
walther@60101
   768
          let val (t, ty) = tm_of vs e
walther@60101
   769
          in case Fdl_Parser.lookup types ty of
walther@60101
   770
              SOME (Fdl_Parser.Array_Type (idxtys, elty)) =>
walther@60101
   771
                let
walther@60101
   772
                  val T = foldr1 HOLogic.mk_prodT
walther@60101
   773
                    (map (mk_type thy prfx) idxtys);
walther@60101
   774
                  val U = mk_type thy prfx elty;
walther@60101
   775
                  val fT = T --> U
walther@60101
   776
                in
walther@60101
   777
                  (Const (\<^const_name>\<open>fun_upd\<close>, fT --> T --> U --> fT) $
walther@60101
   778
                     t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es) $
walther@60101
   779
                       fst (tm_of vs e'),
walther@60101
   780
                   ty)
walther@60101
   781
                end
walther@60101
   782
            | _ => error (ty ^ " is not an array type")
walther@60101
   783
          end
walther@60101
   784
walther@60101
   785
      | tm_of vs (Fdl_Parser.Record (s, flds)) =
walther@60101
   786
          let
walther@60101
   787
            val (T, cmap) = mk_type' thy prfx s;
walther@60101
   788
            val {extension = (ext_name, _), fields, ...} =
walther@60101
   789
              (case get_record_info thy T of
walther@60101
   790
                 NONE => error (s ^ " is not a record type")
walther@60101
   791
               | SOME info => info);
walther@60101
   792
            val flds' = map (apsnd (tm_of vs)) flds;
walther@60101
   793
            val fnames = fields |> invert_map cmap |>
walther@60101
   794
              map (Long_Name.base_name o fst);
walther@60101
   795
            val fnames' = map fst flds;
walther@60101
   796
            val (fvals, ftys) = split_list (map (fn s' =>
walther@60101
   797
              case AList.lookup lcase_eq flds' s' of
walther@60101
   798
                SOME fval_ty => fval_ty
walther@60101
   799
              | NONE => error ("Field " ^ s' ^ " missing in record " ^ s))
walther@60101
   800
                  fnames);
walther@60101
   801
            val _ = (case subtract lcase_eq fnames fnames' of
walther@60101
   802
                [] => ()
walther@60101
   803
              | xs => error ("Extra field(s) " ^ commas xs ^
walther@60101
   804
                  " in record " ^ s));
walther@60101
   805
            val _ = (case duplicates (op =) fnames' of
walther@60101
   806
                [] => ()
walther@60101
   807
              | xs => error ("Duplicate field(s) " ^ commas xs ^
walther@60101
   808
                  " in record " ^ s))
walther@60101
   809
          in
walther@60101
   810
            (list_comb
walther@60101
   811
               (Const (ext_name,
walther@60101
   812
                  map (mk_type thy prfx) ftys @ [HOLogic.unitT] ---> T),
walther@60101
   813
                fvals @ [HOLogic.unit]),
walther@60101
   814
             s)
walther@60101
   815
          end
walther@60101
   816
walther@60101
   817
      | tm_of vs (Fdl_Parser.Array (s, default, assocs)) =
walther@60101
   818
          (case Fdl_Parser.lookup types s of
walther@60101
   819
             SOME (Fdl_Parser.Array_Type (idxtys, elty)) =>
walther@60101
   820
               let
walther@60101
   821
                 val Ts = map (mk_type thy prfx) idxtys;
walther@60101
   822
                 val T = foldr1 HOLogic.mk_prodT Ts;
walther@60101
   823
                 val U = mk_type thy prfx elty;
walther@60101
   824
                 fun mk_idx' T (e, NONE) = HOLogic.mk_set T [fst (tm_of vs e)]
walther@60101
   825
                   | mk_idx' T (e, SOME e') = Const (\<^const_name>\<open>atLeastAtMost\<close>,
walther@60101
   826
                       T --> T --> HOLogic.mk_setT T) $
walther@60101
   827
                         fst (tm_of vs e) $ fst (tm_of vs e');
walther@60101
   828
                 fun mk_idx idx =
walther@60101
   829
                   if length Ts <> length idx then
walther@60101
   830
                     error ("Arity mismatch in construction of array " ^ s)
walther@60101
   831
                   else foldr1 mk_times (map2 mk_idx' Ts idx);
walther@60101
   832
                 fun mk_upd (idxs, e) t =
walther@60101
   833
                   if length idxs = 1 andalso forall (is_none o snd) (hd idxs)
walther@60101
   834
                   then
walther@60101
   835
                     Const (\<^const_name>\<open>fun_upd\<close>, (T --> U) -->
walther@60101
   836
                         T --> U --> T --> U) $ t $
walther@60101
   837
                       foldl1 HOLogic.mk_prod
walther@60101
   838
                         (map (fst o tm_of vs o fst) (hd idxs)) $
walther@60101
   839
                       fst (tm_of vs e)
walther@60101
   840
                   else
walther@60101
   841
                     Const (\<^const_name>\<open>fun_upds\<close>, (T --> U) -->
walther@60101
   842
                         HOLogic.mk_setT T --> U --> T --> U) $ t $
walther@60101
   843
                       foldl1 (HOLogic.mk_binop \<^const_name>\<open>sup\<close>)
walther@60101
   844
                         (map mk_idx idxs) $
walther@60101
   845
                       fst (tm_of vs e)
walther@60101
   846
               in
walther@60101
   847
                 (fold mk_upd assocs
walther@60101
   848
                    (case default of
walther@60101
   849
                       SOME e => Abs ("x", T, fst (tm_of vs e))
walther@60101
   850
                     | NONE => Const (\<^const_name>\<open>undefined\<close>, T --> U)),
walther@60101
   851
                  s)
walther@60101
   852
               end
walther@60101
   853
           | _ => error (s ^ " is not an array type"))
walther@60101
   854
walther@60101
   855
  in tm_of end;
walther@60101
   856
walther@60101
   857
walther@60101
   858
fun mk_rulename (s, i) = Binding.name (s ^ string_of_int i);
walther@60101
   859
walther@60101
   860
walther@60101
   861
fun add_def prfx types pfuns consts (id, (s, e)) (ids as (tab, ctxt), thy) =
walther@60101
   862
  (case Fdl_Parser.lookup consts s of
walther@60101
   863
    SOME ty =>
walther@60101
   864
      let
walther@60101
   865
        val (t, ty') = term_of_expr thy prfx types pfuns ids e;
walther@60101
   866
        val T = mk_type thy prfx ty;
walther@60101
   867
        val T' = mk_type thy prfx ty';
walther@60101
   868
        val _ = T = T' orelse
walther@60101
   869
          error ("Declared type " ^ ty ^ " of " ^ s ^
walther@60101
   870
            "\ndoes not match type " ^ ty' ^ " in definition");
walther@60101
   871
        val id' = mk_rulename id;
walther@60101
   872
        val ((t', (_, th)), lthy') = Named_Target.theory_init thy
walther@60101
   873
          |> Specification.definition NONE [] []
walther@60101
   874
            ((id', []), HOLogic.mk_Trueprop (HOLogic.mk_eq (Free (s, T), t)));
walther@60101
   875
        val phi =
walther@60101
   876
          Proof_Context.export_morphism lthy'
walther@60101
   877
            (Proof_Context.init_global (Proof_Context.theory_of lthy'));
walther@60101
   878
      in
walther@60101
   879
        ((id', Morphism.thm phi th),
walther@60101
   880
          ((Symtab.update (s, (Morphism.term phi t', ty)) tab, Name.declare s ctxt),
walther@60101
   881
            Local_Theory.exit_global lthy'))
walther@60101
   882
       end
walther@60101
   883
  | NONE => error ("Undeclared constant " ^ s));
walther@60101
   884
walther@60101
   885
walther@60101
   886
walther@60101
   887
val add_expr_idents = fold_expr (K I) (insert (op =));
walther@60101
   888
walther@60101
   889
(* sort definitions according to their dependency *)
walther@60101
   890
fun sort_defs _ _ _ _ [] sdefs = rev sdefs
walther@60101
   891
  | sort_defs prfx funs pfuns consts defs sdefs =
walther@60101
   892
      (case find_first (fn (_, (_, e)) =>
walther@60101
   893
         forall (is_some o lookup_prfx prfx pfuns)
walther@60101
   894
           (add_expr_pfuns funs e []) andalso
walther@60101
   895
         forall (fn id =>
walther@60101
   896
           member (fn (s, (_, (s', _))) => s = s') sdefs id orelse
walther@60101
   897
           consts id)
walther@60101
   898
             (add_expr_idents e [])) defs of
walther@60101
   899
         SOME d => sort_defs prfx funs pfuns consts
walther@60101
   900
           (remove (op =) d defs) (d :: sdefs)
walther@60101
   901
       | NONE => error ("Bad definitions: " ^ rulenames defs));
walther@60101
   902
walther@60101
   903
fun add_var prfx (s, ty) (ids, thy) =
walther@60101
   904
  let val ([Free p], ids') = mk_variables thy prfx [s] ty ids
walther@60101
   905
  in (p, (ids', thy)) end;
walther@60101
   906
walther@60101
   907
fun add_init_vars prfx vcs (ids_thy as ((tab, _), _)) =
walther@60101
   908
  fold_map (add_var prfx)
walther@60101
   909
    (map_filter
walther@60101
   910
       (fn s => case try (unsuffix "~") s of
walther@60101
   911
          SOME s' => (case Symtab.lookup tab s' of
walther@60101
   912
            SOME (_, ty) => SOME (s, ty)
walther@60101
   913
          | NONE => error ("Undeclared identifier " ^ s'))
walther@60101
   914
        | NONE => NONE)
walther@60101
   915
       (fold_vcs (add_expr_idents o snd) vcs []))
walther@60101
   916
    ids_thy;
walther@60101
   917
walther@60101
   918
fun term_of_rule thy prfx types pfuns ids rule =
walther@60101
   919
  let val tm_of = fst o term_of_expr thy prfx types pfuns ids
walther@60101
   920
  in case rule of
walther@60101
   921
      Fdl_Parser.Inference_Rule (es, e) => Logic.list_implies
walther@60101
   922
        (map (HOLogic.mk_Trueprop o tm_of) es, HOLogic.mk_Trueprop (tm_of e))
walther@60101
   923
    | Fdl_Parser.Substitution_Rule (es, e, e') => Logic.list_implies
walther@60101
   924
        (map (HOLogic.mk_Trueprop o tm_of) es,
walther@60101
   925
         HOLogic.mk_Trueprop (HOLogic.mk_eq (tm_of e, tm_of e')))
walther@60101
   926
  end;
walther@60101
   927
walther@60101
   928
walther@60101
   929
fun pfun_type thy prfx (argtys, resty) =
walther@60101
   930
  map (mk_type thy prfx) argtys ---> mk_type thy prfx resty;
walther@60101
   931
walther@60101
   932
fun check_pfun_type thy prfx s t optty1 optty2 =
walther@60101
   933
  let
walther@60101
   934
    val T = fastype_of t;
walther@60101
   935
    fun check ty =
walther@60101
   936
      let val U = pfun_type thy prfx ty
walther@60101
   937
      in
walther@60101
   938
        T = U orelse
walther@60101
   939
        error ("Type\n" ^
walther@60101
   940
          Syntax.string_of_typ_global thy T ^
walther@60101
   941
          "\nof function " ^
walther@60101
   942
          Syntax.string_of_term_global thy t ^
walther@60101
   943
          " associated with proof function " ^ s ^
walther@60101
   944
          "\ndoes not match declared type\n" ^
walther@60101
   945
          Syntax.string_of_typ_global thy U)
walther@60101
   946
      end
walther@60101
   947
  in (Option.map check optty1; Option.map check optty2; ()) end;
walther@60101
   948
walther@60101
   949
fun upd_option x y = if is_some x then x else y;
walther@60101
   950
walther@60101
   951
fun check_pfuns_types thy prfx funs =
walther@60101
   952
  Symtab.map (fn s => fn (optty, t) =>
walther@60101
   953
   let val optty' = Fdl_Parser.lookup funs (unprefix_pfun prfx s)
walther@60101
   954
   in
walther@60101
   955
     (check_pfun_type thy prfx s t optty optty';
walther@60101
   956
      (NONE |> upd_option optty |> upd_option optty', t))
walther@60101
   957
   end);
walther@60101
   958
walther@60101
   959
walther@60101
   960
walther@60101
   961
walther@60101
   962
fun pp_vcs msg vcs = Pretty.big_list msg (map (Pretty.str o fst) vcs);
walther@60101
   963
walther@60101
   964
fun pp_open_vcs [] = Pretty.str "All verification conditions have been proved."
walther@60101
   965
  | pp_open_vcs vcs = pp_vcs
walther@60101
   966
      "The following verification conditions remain to be proved:" vcs;
walther@60101
   967
walther@60101
   968
fun partition_vcs vcs = VCtab.fold_rev
walther@60101
   969
  (fn (name, (trace, SOME thms, ps, cs)) =>
walther@60101
   970
        apfst (cons (name, (trace, thms, ps, cs)))
walther@60101
   971
    | (name, (trace, NONE, ps, cs)) =>
walther@60101
   972
        apsnd (cons (name, (trace, ps, cs))))
walther@60101
   973
  vcs ([], []);
walther@60101
   974
walther@60101
   975
fun print_open_vcs f vcs =
walther@60101
   976
  (Pretty.writeln (f (pp_open_vcs (snd (partition_vcs vcs)))); vcs);
walther@60101
   977
walther@60101
   978
fun set_env ctxt defs types funs ids vcs path prefix thy = VCs.map (fn
walther@60101
   979
    {pfuns, type_map, env = NONE} =>
walther@60101
   980
      {pfuns = pfuns,
walther@60101
   981
       type_map = type_map,
walther@60101
   982
       env = SOME
walther@60101
   983
         {ctxt = ctxt, defs = defs, types = types, funs = funs,
walther@60101
   984
          pfuns = check_pfuns_types thy prefix funs pfuns,
walther@60101
   985
          ids = ids, proving = false, vcs = print_open_vcs I vcs, path = path,
walther@60101
   986
          prefix = prefix}}
walther@60101
   987
  | _ => err_unfinished ()) thy;
walther@60101
   988
walther@60101
   989
\<close> ML \<open>
walther@60101
   990
fun set_vcs ({types, vars, consts, funs} : Fdl_Parser.decls)
walther@60101
   991
      (rules, _) ((_, ident), vcs) path opt_prfx thy =
walther@60101
   992
  let
walther@60101
   993
    val prfx' =
walther@60101
   994
      if opt_prfx = "" then
walther@60101
   995
        space_implode "__" (Long_Name.explode (Long_Name.qualifier ident))
walther@60101
   996
      else opt_prfx;
walther@60101
   997
    val prfx = to_lower prfx';
walther@60101
   998
    val {pfuns, ...} = VCs.get thy;
walther@60101
   999
    val (defs, rules') = partition_opt dest_def rules;
walther@60101
  1000
    val consts' =
walther@60101
  1001
      subtract (fn ((_, (s, _)), (s', _)) => s = s') defs (Fdl_Parser.items consts);
walther@60101
  1002
    (* ignore all complex rules in rls files *)
walther@60101
  1003
    val (rules'', other_rules) =
walther@60101
  1004
      List.partition (complex_rule o snd) rules';
walther@60101
  1005
    val _ = if null rules'' then ()
walther@60101
  1006
      else warning ("Ignoring rules: " ^ rulenames rules'');
walther@60101
  1007
walther@60101
  1008
    val vcs' = VCtab.make (maps (fn (tr, vcs) =>
walther@60101
  1009
      map (fn (s, (ps, cs)) => (s, (tr, NONE, ps, cs)))
walther@60101
  1010
        (filter_out (is_trivial_vc o snd) vcs)) vcs);
walther@60101
  1011
walther@60101
  1012
    val _ = (case filter_out (is_some o Fdl_Parser.lookup funs)
walther@60101
  1013
        (pfuns_of_vcs prfx funs pfuns vcs') of
walther@60101
  1014
        [] => ()
walther@60101
  1015
      | fs => error ("Undeclared proof function(s) " ^ commas fs));
walther@60101
  1016
walther@60101
  1017
    val (((defs', vars''), ivars), (ids, thy')) =
walther@60101
  1018
      ((Symtab.empty |>
walther@60101
  1019
        Symtab.update ("false", (\<^term>\<open>False\<close>, booleanN)) |>
walther@60101
  1020
        Symtab.update ("true", (\<^term>\<open>True\<close>, booleanN)),
walther@60101
  1021
        Name.context),
walther@60101
  1022
       thy |> Sign.add_path
walther@60101
  1023
         ((if prfx' = "" then "" else prfx' ^ "__") ^ Long_Name.base_name ident)) |>
walther@60101
  1024
      fold (add_type_def prfx) (Fdl_Parser.items types) |>
walther@60101
  1025
      fold (snd oo add_const prfx) consts' |> (fn ids_thy as ((tab, _), _) =>
walther@60101
  1026
        ids_thy |>
walther@60101
  1027
        fold_map (add_def prfx types pfuns consts)
walther@60101
  1028
          (sort_defs prfx funs pfuns (Symtab.defined tab) defs []) ||>>
walther@60101
  1029
        fold_map (add_var prfx) (Fdl_Parser.items vars) ||>>
walther@60101
  1030
        add_init_vars prfx vcs');
walther@60101
  1031
walther@60101
  1032
    val ctxt =
walther@60101
  1033
      [Element.Fixes (map (fn (s, T) =>
walther@60101
  1034
         (Binding.name s, SOME T, NoSyn)) (vars'' @ ivars)),
walther@60101
  1035
       Element.Assumes (map (fn (id, rl) =>
walther@60101
  1036
         ((mk_rulename id, []),
walther@60101
  1037
          [(term_of_rule thy' prfx types pfuns ids rl, [])]))
walther@60101
  1038
           other_rules),
walther@60101
  1039
       Element.Notes ("", [((Binding.name "defns", []), map (rpair [] o single o snd) defs')])]
walther@60101
  1040
walther@60101
  1041
  in
walther@60101
  1042
    set_env ctxt defs' types funs ids vcs' path prfx thy'
walther@60101
  1043
  end;
walther@60097
  1044
\<close> ML \<open>
walther@60097
  1045
\<close>
walther@60097
  1046
walther@60101
  1047
subsection \<open>cp. HOL/SPARK/Tools/spark_commands.ML\<close>
walther@60097
  1048
ML \<open>
walther@60097
  1049
\<close> ML \<open>
walther@60101
  1050
(*  Title:      src/Tools/isac/BridgeJEdit/cp_spark_commands.ML
walther@60101
  1051
    Author:     Walther Neuper, JKU Linz
walther@60101
  1052
    (c) due to copyright terms
walther@60101
  1053
walther@60101
  1054
Preliminary code for developing Outer_Syntax.command..Example from spark_open as a model.
walther@60101
  1055
*)
walther@60101
  1056
walther@60101
  1057
fun spark_open           header (files, prfx) thy =
walther@60101
  1058
(*Fdl_Lexer.vcg_header ..^^^^^^ ^^^^^^^^^^^^^..from Resources.provide_parse_files: are all funs!*)
walther@60101
  1059
  let
walther@60101
  1060
(**)val _ = @{print} {a = "//--- ###C spark_open", files = files, header = header, prfx = prfx};(**)
walther@60101
  1061
    val xxx as ([{src_path, lines = vc_lines, pos = vc_pos, ...}: Token.file, (*.siv*)
walther@60101
  1062
(**)  {lines = fdl_lines, pos = fdl_pos, ...},                                (*.fld*)
walther@60101
  1063
(**)  {lines = rls_lines, pos = rls_pos, ...}                                 (*.rls*)
walther@60101
  1064
      ], thy') = 
walther@60101
  1065
    (@{print} {a = "###C spark_open: before call <files thy>, files = siv_header !"}; files thy);
walther@60101
  1066
(*                                                                                    ^^^^^ arg!*)
walther@60101
  1067
    val path = fst (Path.split_ext src_path);
walther@60101
  1068
  in
walther@60101
  1069
    @{print} {a = "###C spark_open: result from siv_header, without thy'", fst_xxx = fst xxx};
walther@60101
  1070
    @{print} {a = "###C spark_open: 1 arg for SPARK_VCs.set_vcs",
walther@60101
  1071
(** )  arg1 = (snd (Fdl_Parser.parse_declarations fdl_pos (cat_lines fdl_lines))),
walther@60101
  1072
      arg2 = (Fdl_Parser.parse_rules rls_pos (cat_lines rls_lines)),            ( **)
walther@60101
  1073
      arg3 = (snd (Fdl_Parser.parse_vcs header vc_pos (cat_lines vc_lines))),
walther@60101
  1074
(*                                                     !^^^^^^^!*)
walther@60101
  1075
      z = "\\--- ###C spark_open"};(**)
walther@60101
  1076
    (*SPARK_VCs.*)set_vcs
walther@60101
  1077
(**)  (snd (Fdl_Parser.parse_declarations fdl_pos (cat_lines fdl_lines)))(** )Fdl_Parser.empty_decls( **)
walther@60101
  1078
(**)  (Fdl_Parser.parse_rules rls_pos (cat_lines rls_lines))             (** )Fdl_Parser.empty_rules( **)
walther@60101
  1079
      (snd (Fdl_Parser.parse_vcs header vc_pos (cat_lines vc_lines)))
walther@60101
  1080
      path prfx thy'
walther@60101
  1081
  end;
walther@60097
  1082
\<close> ML \<open>
walther@60097
  1083
\<close>
walther@60097
  1084
walther@60102
  1085
section \<open>Adapt Fdl_Lexer.scan -> parse Formalise.T\<close>
walther@60102
  1086
text \<open>
walther@60102
  1087
  Possible models, where a file is read and parsed:
walther@60102
  1088
  * for *.bib see usage of Resources.provide_parse_files     .. parsed by LaTeX
walther@60102
  1089
  * Outer_Syntax.command \ <command_keyword>\<open>external_file\<close>  .. ?!?
walther@60102
  1090
  * Outer_Syntax.command \<command_keyword>\<open>spark_open\<close>
walther@60102
  1091
  We pursue the latter.
walther@60102
  1092
  From c.f.2b9b4b38ab96 we have found Fdl_Lexer.scan as the very model
walther@60102
  1093
  for scanning/parsing Formalise.T.
walther@60102
  1094
\<close>
walther@60102
  1095
subsection \<open>testbed separating Fdl_Lexer.scan from SPARK code\<close>
walther@60102
  1096
ML_file "~~/test/HOL/SPARK/test-data.sml"  ML \<open>(*..:*)g_c_d_siv_content\<close>
walther@60102
  1097
declare [[ML_print_depth = 10]] (* 7 9999999999 *)
walther@60102
  1098
ML \<open>
walther@60102
  1099
open Fdl_Lexer
walther@60102
  1100
(*
walther@60102
  1101
  step 1: go top down and keep respective results according to trace from
walther@60102
  1102
          spark_open \<open>greatest_common_divisor/g_c_d\<close> SPARK/../Greatest_Common_Divisor.thy.
walther@60102
  1103
          (in the example this is simple: just watch the bottom of Output
walther@60102
  1104
           The following verification conditions remain to be proved:
walther@60102
  1105
             procedure_g_c_d_4
walther@60102
  1106
             procedure_g_c_d_11)
walther@60102
  1107
        ! remove @~{print} from calling functions (higher level) 
walther@60102
  1108
          once you go down to the called functions (lower level)
walther@60102
  1109
          in order to reduce noise and to focus the code under actual consideration.
walther@60102
  1110
  step 2: at bottom gather the most elementary funs for transformation
walther@60102
  1111
           of g_c_d_siv_content to args of Fdl_Lexer.scan
walther@60102
  1112
*)
walther@60102
  1113
\<close> ML \<open>
walther@60102
  1114
\<close> ML \<open> (** high start level: parsing to header + VCs **)
walther@60102
  1115
val parsed = Fdl_Parser.parse_vcs Fdl_Lexer.siv_header Position.start g_c_d_siv_content;
walther@60102
  1116
\<close> ML \<open> (*               ^^^^^^^^^ <<<<<------------------------------------------------------*)
walther@60102
  1117
val ((banner, (date, time), (date2, time2), (string, string2)), ((str, str2), vcss)) = parsed;
walther@60102
  1118
walther@60102
  1119
if banner = ("Semantic Analysis of SPARK Text",
walther@60102
  1120
    "Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039",
walther@60102
  1121
    "Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.")  andalso
walther@60102
  1122
  (date, time) = (("29", "NOV", "2010"), ("14", "30", "10", NONE)) andalso
walther@60102
  1123
  (date2, time2) = (("29", "NOV", "2010"), ("14", "30", "11", NONE)) andalso
walther@60102
  1124
  (string, string2) =
walther@60102
  1125
    ("SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039",
walther@60102
  1126
    "Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.") andalso
walther@60102
  1127
  (str, str2) = ("procedure", "Greatest_Common_Divisor.G_C_D")
walther@60102
  1128
  then () else error "Fdl_Parser.parse_vcs header CHANGED";
walther@60102
  1129
if length vcss = 11 (* verification conditions (VCs) from g_c_d_siv_content *) then
walther@60102
  1130
  case vcss of
walther@60102
  1131
    ("path(s) from start to run-time check associated with statement of line 8",
walther@60102
  1132
      [("procedure_g_c_d_1", ([], [("1", _(*Ident "true": Fdl_Parser.expr*))]))]) :: _ => ()
walther@60102
  1133
  | _ => error "Fdl_Parser.parse_vcs vcss CHANGED 1"
walther@60102
  1134
else error "Fdl_Parser.parse_vcs vcss CHANGED 2";
walther@60102
  1135
\<close> ML \<open>
walther@60102
  1136
\<close> ML \<open>
walther@60102
  1137
\<close> ML \<open> (* lower level 1: tokenize string to header + VCs *)
walther@60102
  1138
val level1_header_toks =
walther@60102
  1139
(* tokenize to header + VCs is NON-standard ..
walther@60102
  1140
            vvvvvvvv--- parses vvvvvvvvvv     +     vvvvvvvvv, but leaves vcsss as Token list*)
walther@60102
  1141
  Fdl_Lexer.tokenize Fdl_Lexer.siv_header Fdl_Lexer.c_comment Position.start g_c_d_siv_content;
walther@60102
  1142
\<close> ML \<open> (*   ^^^^^^^^ <<<<<-------------------------------------------------------------------*)
walther@60102
  1143
val ((banner, (date, time), (date2, time2), (string, string2)), toks ) = level1_header_toks;
walther@60102
  1144
walther@60102
  1145
if banner = ("Semantic Analysis of SPARK Text",
walther@60102
  1146
    "Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039",
walther@60102
  1147
    "Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.")  andalso
walther@60102
  1148
  (date, time) = (("29", "NOV", "2010"), ("14", "30", "10", NONE)) andalso
walther@60102
  1149
  (date2, time2) = (("29", "NOV", "2010"), ("14", "30", "11", NONE)) andalso
walther@60102
  1150
  (string, string2) =
walther@60102
  1151
    ("SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039",
walther@60102
  1152
    "Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.")
walther@60102
  1153
  then () else error "Fdl_Lexer.tokenize header CHANGED";
walther@60102
  1154
if length toks = 319 (* tokens for str :: str :: verification conditions from g_c_d_siv_content *) then
walther@60102
  1155
  case toks of
walther@60102
  1156
    Token (Keyword, "procedure", _(*Position.T*)) :: 
walther@60102
  1157
    Token (Long_Ident, "Greatest_Common_Divisor.G_C_D", _(*Position.T*)) :: _  => ()
walther@60102
  1158
  | _ => error "Fdl_Lexer.tokenize tokens CHANGED 1"
walther@60102
  1159
else error "Fdl_Lexer.tokenize header CHANGED 2";
walther@60102
  1160
\<close> ML \<open>
walther@60102
  1161
\<close> ML \<open>
walther@60102
  1162
\<close> ML \<open> (* lower level 2: lexing the string *)
walther@60102
  1163
val chars = Fdl_Lexer.explode_pos g_c_d_siv_content(*_content*) Position.start;
walther@60102
  1164
if length chars = 2886 then () else error "Fdl_Lexer.explode_pos g_c_d_siv_content CHANGED";
walther@60102
  1165
val level2_header_toks = (Scan.finite char_stopper
walther@60102
  1166
  (Scan.error (Fdl_Lexer.scan Fdl_Lexer.siv_header Fdl_Lexer.c_comment)) chars);
walther@60102
  1167
(* ---------------------------------------------------------------------------\
walther@60102
  1168
               Fdl_Lexer.scan Fdl_Lexer.siv_header Fdl_Lexer.c_comment chars
walther@60102
  1169
exception MORE () raised (line 176 of "General/scan.ML") ---------------------/ *)
walther@60102
  1170
walther@60102
  1171
case level2_header_toks of
walther@60102
  1172
  (((("Semantic Analysis of SPARK Text", 
walther@60102
  1173
       "Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039", 
walther@60102
  1174
       "Copyright (C) 2010 Altran Praxis Limited, Bath, U.K."), 
walther@60102
  1175
       (("29", "NOV", "2010"), ("14", "30", "10", NONE)),
walther@60102
  1176
       (("29", "NOV", "2010"), ("14", "30", "11", NONE)), 
walther@60102
  1177
       ("SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039", 
walther@60102
  1178
         "Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.")),
walther@60102
  1179
     Token (Keyword, "procedure", _) :: Token (Long_Ident, "Greatest_Common_Divisor.G_C_D", _) ::
walther@60102
  1180
     Token (Traceability, 
walther@60102
  1181
       "path(s) from start to run-time check associated with statement of line 8", _) ::
walther@60102
  1182
     _ ),
walther@60102
  1183
   []) => ()
walther@60102
  1184
| _ => error "Fdl_Lexer.scan level2 CHANGED";
walther@60102
  1185
\<close> ML \<open>
walther@60102
  1186
\<close>
walther@60102
  1187
subsection \<open>adapt testbed towards Isac's Formalise.T parser\<close>
walther@60102
  1188
text \<open>
walther@60102
  1189
\<close> ML \<open>
walther@60102
  1190
\<close> ML \<open>
walther@60102
  1191
\<close> ML \<open>
walther@60102
  1192
\<close>
walther@60102
  1193
walther@60101
  1194
walther@60097
  1195
section \<open>setup command_keyword \<close>
walther@60097
  1196
subsection \<open>hide for development, use for Test_Isac\<close>
walther@60097
  1197
ML \<open>
walther@60101
  1198
(*HIDE for development, activate for Test_Isac*)
walther@60095
  1199
type chars = (string * Position.T) list;
walther@60095
  1200
val e_chars = []: (string * Position.T) list;
walther@60095
  1201
walther@60096
  1202
fun spark_open (_: ((*Fdl_Lexer.*)chars -> 'a * (*Fdl_Lexer.*)chars)) 
walther@60097
  1203
    (_: ('b -> Token.file list * theory) * string) (_: 'b) = @{theory}
walther@60096
  1204
walther@60095
  1205
type banner = string * string * string;               val e_banner = ("b1", "b2", "b3")
walther@60095
  1206
type date = string * string * string;                 val e_date = ("d1", "d2", "d3")
walther@60095
  1207
type time = string * string * string * string option; val e_time = ("t1", "t2", "t3", SOME "t4")
walther@60095
  1208
walther@60096
  1209
fun siv_header (_: chars) = 
walther@60096
  1210
  ((e_banner, (e_date, e_time), (e_date, e_time), ("", "")), e_chars);
walther@60101
  1211
(*HIDE for development, activate for Test_Isac*)
walther@60095
  1212
\<close> ML \<open>
walther@60097
  1213
\<close> 
walther@60097
  1214
subsection \<open>def.for Build_Isac; for Test_Isac outcomment SPARK_Commands, Fdl_Lexer below\<close>
walther@60097
  1215
ML \<open>
walther@60095
  1216
\<close> ML \<open>
walther@60097
  1217
val _ =                                                  
walther@60097
  1218
  Outer_Syntax.command \<^command_keyword>\<open>Example\<close> "start Isac Calculation from a Formalise-file"
walther@60097
  1219
    (*preliminary from spark_open*)
walther@60097
  1220
    (Resources.provide_parse_files "spark_open" -- Parse.parname
walther@60097
  1221
(** )  >> (Toplevel.theory o ((*SPARK_Commands.*)spark_open Fdl_Lexer.siv_header)));( **)
walther@60097
  1222
(**)  >> (Toplevel.theory o ((*SPARK_Commands.*)spark_open (*Fdl_Lexer.*)siv_header)));(**)
walther@60095
  1223
\<close> ML \<open>
walther@60095
  1224
\<close>
walther@60095
  1225
walther@60096
  1226
end