src/HOL/SPARK/Tools/spark_vcs.ML
author Andreas Lochbihler
Wed, 08 Apr 2015 14:59:02 +0200
changeset 59954 5ee7e9721eac
parent 59810 e749a0f2f401
child 60754 02924903a6fd
permissions -rw-r--r--
more lemmas and operations on cset (adapted from FSet)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/SPARK/Tools/spark_vcs.ML
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     3
    Copyright:  secunet Security Networks AG
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     4
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     5
Store for verification conditions generated by SPARK/Ada.
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     6
*)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     7
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     8
signature SPARK_VCS =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     9
sig
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    10
  val set_vcs: Fdl_Parser.decls -> Fdl_Parser.rules ->
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    11
    (string * string) * Fdl_Parser.vcs -> Path.T -> string -> theory -> theory
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    12
  val add_proof_fun: (typ option -> 'a -> term) ->
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    13
    string * ((string list * string) option * 'a) ->
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    14
    theory -> theory
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
    15
  val add_type: string * (typ * (string * string) list) -> theory -> theory
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
    16
  val lookup_vc: theory -> bool -> string -> (Element.context_i list *
41896
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41878
diff changeset
    17
    (string * thm list option * Element.context_i * Element.statement_i)) option
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
    18
  val get_vcs: theory -> bool ->
41896
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41878
diff changeset
    19
    Element.context_i list * (binding * thm) list * (string *
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41878
diff changeset
    20
    (string * thm list option * Element.context_i * Element.statement_i)) list
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41878
diff changeset
    21
  val mark_proved: string -> thm list -> theory -> theory
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
    22
  val close: bool -> theory -> theory
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    23
  val is_closed: theory -> bool
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    24
end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    25
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    26
structure SPARK_VCs: SPARK_VCS =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    27
struct
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    28
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    29
open Fdl_Parser;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    30
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    31
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    32
(** theory data **)
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    33
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    34
fun err_unfinished () = error "An unfinished SPARK environment is still open."
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    35
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58956
diff changeset
    36
val strip_number = apply2 implode o take_suffix Fdl_Lexer.is_digit o raw_explode;
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    37
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    38
val name_ord = prod_ord string_ord (option_ord int_ord) o
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58956
diff changeset
    39
  apply2 (strip_number ##> Int.fromString);
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    40
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    41
structure VCtab = Table(type key = string val ord = name_ord);
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    42
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    43
structure VCs = Theory_Data
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    44
(
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    45
  type T =
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    46
    {pfuns: ((string list * string) option * term) Symtab.table,
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
    47
     type_map: (typ * (string * string) list) Symtab.table,
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    48
     env:
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    49
       {ctxt: Element.context_i list,
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    50
        defs: (binding * thm) list,
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    51
        types: fdl_type tab,
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    52
        funs: (string list * string) tab,
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
    53
        pfuns: ((string list * string) option * term) Symtab.table,
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    54
        ids: (term * string) Symtab.table * Name.context,
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    55
        proving: bool,
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    56
        vcs: (string * thm list option *
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    57
          (string * expr) list * (string * expr) list) VCtab.table,
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    58
        path: Path.T,
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    59
        prefix: string} option}
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    60
  val empty : T = {pfuns = Symtab.empty, type_map = Symtab.empty, env = NONE}
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    61
  val extend = I
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    62
  fun merge ({pfuns = pfuns1, type_map = type_map1, env = NONE},
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    63
        {pfuns = pfuns2, type_map = type_map2, env = NONE}) =
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    64
        {pfuns = Symtab.merge (eq_pair (op =) (op aconv)) (pfuns1, pfuns2),
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    65
         type_map = Symtab.merge (op =) (type_map1, type_map2),
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    66
         env = NONE}
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    67
    | merge _ = err_unfinished ()
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    68
)
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    69
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    70
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    71
(** utilities **)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    72
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    73
val to_lower = raw_explode #> map Symbol.to_ascii_lower #> implode;
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    74
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58956
diff changeset
    75
val lcase_eq = (op =) o apply2 (to_lower o Long_Name.base_name);
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
    76
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    77
fun lookup_prfx "" tab s = Symtab.lookup tab s
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    78
  | lookup_prfx prfx tab s = (case Symtab.lookup tab s of
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    79
        NONE => Symtab.lookup tab (prfx ^ "__" ^ s)
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    80
      | x => x);
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    81
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    82
fun strip_prfx s =
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    83
  let
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    84
    fun strip ys [] = ("", implode ys)
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    85
      | strip ys ("_" :: "_" :: xs) = (implode (rev xs), implode ys)
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    86
      | strip ys (x :: xs) = strip (x :: ys) xs
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    87
  in strip [] (rev (raw_explode s)) end;
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    88
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
    89
fun unprefix_pfun "" s = s
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
    90
  | unprefix_pfun prfx s =
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
    91
      let val (prfx', s') = strip_prfx s
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
    92
      in if prfx = prfx' then s' else s end;
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
    93
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    94
fun mk_unop s t =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    95
  let val T = fastype_of t
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    96
  in Const (s, T --> T) $ t end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    97
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    98
fun mk_times (t, u) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    99
  let
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   100
    val setT = fastype_of t;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   101
    val T = HOLogic.dest_setT setT;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   102
    val U = HOLogic.dest_setT (fastype_of u)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   103
  in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   104
    Const (@{const_name Sigma}, setT --> (T --> HOLogic.mk_setT U) -->
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   105
      HOLogic.mk_setT (HOLogic.mk_prodT (T, U))) $ t $ Abs ("", T, u)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   106
  end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   107
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   108
fun get_type thy prfx ty =
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   109
  let val {type_map, ...} = VCs.get thy
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   110
  in lookup_prfx prfx type_map ty end;
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   111
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   112
fun mk_type' _ _ "integer" = (HOLogic.intT, [])
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   113
  | mk_type' _ _ "boolean" = (HOLogic.boolT, [])
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   114
  | mk_type' thy prfx ty =
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   115
      (case get_type thy prfx ty of
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   116
         NONE =>
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   117
           (Syntax.check_typ (Proof_Context.init_global thy)
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   118
              (Type (Sign.full_name thy (Binding.name ty), [])),
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   119
            [])
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   120
       | SOME p => p);
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   121
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   122
fun mk_type thy prfx ty = fst (mk_type' thy prfx ty);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   123
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   124
val booleanN = "boolean";
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   125
val integerN = "integer";
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   126
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   127
fun define_overloaded (def_name, eq) lthy =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   128
  let
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   129
    val ((c, _), rhs) = eq |> Syntax.check_term lthy |>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   130
      Logic.dest_equals |>> dest_Free;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   131
    val ((_, (_, thm)), lthy') = Local_Theory.define
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   132
      ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs)) lthy
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42356
diff changeset
   133
    val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy');
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42356
diff changeset
   134
    val thm' = singleton (Proof_Context.export lthy' ctxt_thy) thm
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   135
  in (thm', lthy') end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   136
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   137
fun strip_underscores s =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   138
  strip_underscores (unsuffix "_" s) handle Fail _ => s;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   139
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   140
fun strip_tilde s =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   141
  unsuffix "~" s ^ "_init" handle Fail _ => s;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   142
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   143
val mangle_name = strip_underscores #> strip_tilde;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   144
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   145
fun mk_variables thy prfx xs ty (tab, ctxt) =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   146
  let
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   147
    val T = mk_type thy prfx ty;
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42499
diff changeset
   148
    val (ys, ctxt') = fold_map Name.variant (map mangle_name xs) ctxt;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   149
    val zs = map (Free o rpair T) ys;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   150
  in (zs, (fold (Symtab.update o apsnd (rpair ty)) (xs ~~ zs) tab, ctxt')) end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   151
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   152
fun get_record_info thy T = (case Record.dest_recTs T of
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   153
    [(tyname, [@{typ unit}])] =>
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   154
      Record.get_info thy (Long_Name.qualifier tyname)
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   155
  | _ => NONE);
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   156
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   157
fun find_field [] fname fields =
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   158
      find_first (curry lcase_eq fname o fst) fields
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   159
  | find_field cmap fname fields =
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   160
      (case AList.lookup (op =) cmap fname of
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   161
         NONE => NONE
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   162
       | SOME fname' => SOME (fname', the (AList.lookup (op =) fields fname')));
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   163
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   164
fun find_field' fname = get_first (fn (flds, fldty) =>
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   165
  if member (op =) flds fname then SOME fldty else NONE);
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   166
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   167
fun assoc_ty_err thy T s msg =
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   168
  error ("Type " ^ Syntax.string_of_typ_global thy T ^
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   169
    " associated with SPARK type " ^ s ^ "\n" ^ msg);
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   170
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   171
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   172
(** generate properties of enumeration types **)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   173
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   174
fun add_enum_type tyname tyname' thy =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   175
  let
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58130
diff changeset
   176
    val {case_name, ...} = the (BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] tyname');
58130
5e9170812356 ported to use new-style datatypes
blanchet
parents: 58112
diff changeset
   177
    val cs = map Const (the (BNF_LFP_Compat.get_constrs thy tyname'));
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   178
    val k = length cs;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   179
    val T = Type (tyname', []);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   180
    val p = Const (@{const_name pos}, T --> HOLogic.intT);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   181
    val v = Const (@{const_name val}, HOLogic.intT --> T);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   182
    val card = Const (@{const_name card},
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   183
      HOLogic.mk_setT T --> HOLogic.natT) $ HOLogic.mk_UNIV T;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   184
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   185
    fun mk_binrel_def s f = Logic.mk_equals
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   186
      (Const (s, T --> T --> HOLogic.boolT),
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   187
       Abs ("x", T, Abs ("y", T,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   188
         Const (s, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   189
           (f $ Bound 1) $ (f $ Bound 0))));
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   190
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   191
    val (((def1, def2), def3), lthy) = thy |>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   192
42416
a8a9f4d79196 - renamed enum type class to spark_enum, to avoid confusion with
berghofe
parents: 42396
diff changeset
   193
      Class.instantiation ([tyname'], [], @{sort spark_enum}) |>
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   194
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   195
      define_overloaded ("pos_" ^ tyname ^ "_def", Logic.mk_equals
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   196
        (p,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   197
         list_comb (Const (case_name, replicate k HOLogic.intT @
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   198
             [T] ---> HOLogic.intT),
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   199
           map (HOLogic.mk_number HOLogic.intT) (0 upto k - 1)))) ||>>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   200
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   201
      define_overloaded ("less_eq_" ^ tyname ^ "_def",
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   202
        mk_binrel_def @{const_name less_eq} p) ||>>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   203
      define_overloaded ("less_" ^ tyname ^ "_def",
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   204
        mk_binrel_def @{const_name less} p);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   205
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   206
    val UNIV_eq = Goal.prove lthy [] []
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   207
      (HOLogic.mk_Trueprop (HOLogic.mk_eq
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   208
         (HOLogic.mk_UNIV T, HOLogic.mk_set T cs)))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   209
      (fn _ =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   210
         rtac @{thm subset_antisym} 1 THEN
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   211
         rtac @{thm subsetI} 1 THEN
58956
a816aa3ff391 proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
wenzelm
parents: 58354
diff changeset
   212
         Old_Datatype_Aux.exh_tac lthy (K (#exhaust (BNF_LFP_Compat.the_info
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58130
diff changeset
   213
           (Proof_Context.theory_of lthy) [BNF_LFP_Compat.Keep_Nesting] tyname'))) 1 THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   214
         ALLGOALS (asm_full_simp_tac lthy));
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   215
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   216
    val finite_UNIV = Goal.prove lthy [] []
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   217
      (HOLogic.mk_Trueprop (Const (@{const_name finite},
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   218
         HOLogic.mk_setT T --> HOLogic.boolT) $ HOLogic.mk_UNIV T))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   219
      (fn _ => simp_tac (lthy addsimps [UNIV_eq]) 1);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   220
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   221
    val card_UNIV = Goal.prove lthy [] []
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   222
      (HOLogic.mk_Trueprop (HOLogic.mk_eq
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   223
         (card, HOLogic.mk_number HOLogic.natT k)))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   224
      (fn _ => simp_tac (lthy addsimps [UNIV_eq]) 1);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   225
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   226
    val range_pos = Goal.prove lthy [] []
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   227
      (HOLogic.mk_Trueprop (HOLogic.mk_eq
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   228
         (Const (@{const_name image}, (T --> HOLogic.intT) -->
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   229
            HOLogic.mk_setT T --> HOLogic.mk_setT HOLogic.intT) $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   230
              p $ HOLogic.mk_UNIV T,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   231
          Const (@{const_name atLeastLessThan}, HOLogic.intT -->
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   232
            HOLogic.intT --> HOLogic.mk_setT HOLogic.intT) $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   233
              HOLogic.mk_number HOLogic.intT 0 $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   234
              (@{term int} $ card))))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   235
      (fn _ =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   236
         simp_tac (lthy addsimps [card_UNIV]) 1 THEN
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   237
         simp_tac (lthy addsimps [UNIV_eq, def1]) 1 THEN
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   238
         rtac @{thm subset_antisym} 1 THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   239
         simp_tac lthy 1 THEN
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   240
         rtac @{thm subsetI} 1 THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   241
         asm_full_simp_tac (lthy addsimps @{thms interval_expand}
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   242
           delsimps @{thms atLeastLessThan_iff}) 1);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   243
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   244
    val lthy' =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   245
      Class.prove_instantiation_instance (fn ctxt =>
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   246
        Class.intro_classes_tac ctxt [] THEN
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   247
        rtac finite_UNIV 1 THEN
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   248
        rtac range_pos 1 THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   249
        simp_tac (put_simpset HOL_basic_ss ctxt addsimps [def3]) 1 THEN
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   250
        simp_tac (put_simpset HOL_basic_ss ctxt addsimps [def2]) 1) lthy;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   251
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   252
    val (pos_eqs, val_eqs) = split_list (map_index (fn (i, c) =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   253
      let
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   254
        val n = HOLogic.mk_number HOLogic.intT i;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   255
        val th = Goal.prove lthy' [] []
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   256
          (HOLogic.mk_Trueprop (HOLogic.mk_eq (p $ c, n)))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   257
          (fn _ => simp_tac (lthy' addsimps [def1]) 1);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   258
        val th' = Goal.prove lthy' [] []
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   259
          (HOLogic.mk_Trueprop (HOLogic.mk_eq (v $ n, c)))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   260
          (fn _ =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   261
             rtac (@{thm inj_pos} RS @{thm injD}) 1 THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   262
             simp_tac (lthy' addsimps [@{thm pos_val}, range_pos, card_UNIV, th]) 1)
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   263
      in (th, th') end) cs);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   264
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   265
    val first_el = Goal.prove lthy' [] []
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   266
      (HOLogic.mk_Trueprop (HOLogic.mk_eq
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   267
         (Const (@{const_name first_el}, T), hd cs)))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   268
      (fn _ => simp_tac (lthy' addsimps [@{thm first_el_def}, hd val_eqs]) 1);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   269
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   270
    val last_el = Goal.prove lthy' [] []
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   271
      (HOLogic.mk_Trueprop (HOLogic.mk_eq
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   272
         (Const (@{const_name last_el}, T), List.last cs)))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50126
diff changeset
   273
      (fn _ => simp_tac (lthy' addsimps [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   274
  in
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   275
    lthy' |>
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   276
    Local_Theory.note
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45585
diff changeset
   277
      ((Binding.name (tyname ^ "_card"), @{attributes [simp]}), [card_UNIV]) ||>>
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   278
    Local_Theory.note
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45585
diff changeset
   279
      ((Binding.name (tyname ^ "_pos"), @{attributes [simp]}), pos_eqs) ||>>
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   280
    Local_Theory.note
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45585
diff changeset
   281
      ((Binding.name (tyname ^ "_val"), @{attributes [simp]}), val_eqs) ||>>
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   282
    Local_Theory.note
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45585
diff changeset
   283
      ((Binding.name (tyname ^ "_first_el"), @{attributes [simp]}), [first_el]) ||>>
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   284
    Local_Theory.note
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45585
diff changeset
   285
      ((Binding.name (tyname ^ "_last_el"), @{attributes [simp]}), [last_el]) |> snd |>
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   286
    Local_Theory.exit_global
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   287
  end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   288
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   289
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   290
fun check_no_assoc thy prfx s = case get_type thy prfx s of
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   291
    NONE => ()
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   292
  | SOME _ => error ("Cannot associate a type with " ^ s ^
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   293
      "\nsince it is no record or enumeration type");
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   294
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   295
fun check_enum [] [] = NONE 
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   296
  | check_enum els [] = SOME ("has no element(s) " ^ commas els)
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   297
  | check_enum [] cs = SOME ("has extra element(s) " ^
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   298
      commas (map (Long_Name.base_name o fst) cs))
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   299
  | check_enum (el :: els) ((cname, _) :: cs) =
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   300
      if lcase_eq (el, cname) then check_enum els cs
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   301
      else SOME ("either has no element " ^ el ^
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   302
        " or it is at the wrong position");
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   303
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   304
fun invert_map [] = I
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   305
  | invert_map cmap =
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   306
      map (apfst (the o AList.lookup (op =) (map swap cmap)));
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   307
 
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   308
fun add_type_def prfx (s, Basic_Type ty) (ids, thy) =
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   309
      (check_no_assoc thy prfx s;
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   310
       (ids,
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   311
        Typedecl.abbrev_global (Binding.name s, [], NoSyn)
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   312
          (mk_type thy prfx ty) thy |> snd))
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   313
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   314
  | add_type_def prfx (s, Enum_Type els) ((tab, ctxt), thy) =
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   315
      let
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   316
        val (thy', tyname) = (case get_type thy prfx s of
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   317
            NONE =>
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   318
              let
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   319
                val tyb = Binding.name s;
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   320
                val tyname = Sign.full_name thy tyb
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   321
              in
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   322
                (thy |>
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58130
diff changeset
   323
                 BNF_LFP_Compat.add_datatype [BNF_LFP_Compat.Keep_Nesting]
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   324
                   [((tyb, [], NoSyn),
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   325
                     map (fn s => (Binding.name s, [], NoSyn)) els)] |> snd |>
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   326
                 add_enum_type s tyname,
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   327
                 tyname)
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   328
              end
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   329
          | SOME (T as Type (tyname, []), cmap) =>
58130
5e9170812356 ported to use new-style datatypes
blanchet
parents: 58112
diff changeset
   330
              (case BNF_LFP_Compat.get_constrs thy tyname of
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   331
                 NONE => assoc_ty_err thy T s "is not a datatype"
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   332
               | SOME cs =>
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   333
                   let val (prfx', _) = strip_prfx s
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   334
                   in
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   335
                     case check_enum (map (unprefix_pfun prfx') els)
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   336
                         (invert_map cmap cs) of
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   337
                       NONE => (thy, tyname)
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   338
                     | SOME msg => assoc_ty_err thy T s msg
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   339
                   end)
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   340
          | SOME (T, _) => assoc_ty_err thy T s "is not a datatype");
58130
5e9170812356 ported to use new-style datatypes
blanchet
parents: 58112
diff changeset
   341
        val cs = map Const (the (BNF_LFP_Compat.get_constrs thy' tyname));
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   342
      in
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   343
        ((fold (Symtab.update_new o apsnd (rpair s)) (els ~~ cs) tab,
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   344
          fold Name.declare els ctxt),
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   345
         thy')
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   346
      end
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   347
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   348
  | add_type_def prfx (s, Array_Type (argtys, resty)) (ids, thy) =
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   349
      (check_no_assoc thy prfx s;
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   350
       (ids,
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   351
        Typedecl.abbrev_global (Binding.name s, [], NoSyn)
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   352
          (foldr1 HOLogic.mk_prodT (map (mk_type thy prfx) argtys) -->
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   353
             mk_type thy prfx resty) thy |> snd))
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   354
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   355
  | add_type_def prfx (s, Record_Type fldtys) (ids, thy) =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   356
      (ids,
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   357
       let val fldTs = maps (fn (flds, ty) =>
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   358
         map (rpair (mk_type thy prfx ty)) flds) fldtys
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   359
       in case get_type thy prfx s of
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   360
           NONE =>
44653
6d8d09b90398 discontinued slightly odd "Defining record ..." message and corresponding quiet_mode;
wenzelm
parents: 44030
diff changeset
   361
             Record.add_record ([], Binding.name s) NONE
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   362
               (map (fn (fld, T) => (Binding.name fld, T, NoSyn)) fldTs) thy
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   363
         | SOME (rT, cmap) =>
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   364
             (case get_record_info thy rT of
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   365
                NONE => assoc_ty_err thy rT s "is not a record type"
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   366
              | SOME {fields, ...} =>
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   367
                  let val fields' = invert_map cmap fields
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   368
                  in
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58956
diff changeset
   369
                    (case subtract (lcase_eq o apply2 fst) fldTs fields' of
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   370
                       [] => ()
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   371
                     | flds => assoc_ty_err thy rT s ("has extra field(s) " ^
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   372
                         commas (map (Long_Name.base_name o fst) flds));
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   373
                     map (fn (fld, T) =>
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   374
                       case AList.lookup lcase_eq fields' fld of
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   375
                         NONE => assoc_ty_err thy rT s ("has no field " ^ fld)
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   376
                       | SOME U => T = U orelse assoc_ty_err thy rT s
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   377
                           ("has field " ^
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   378
                            fld ^ " whose type\n" ^
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   379
                            Syntax.string_of_typ_global thy U ^
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   380
                            "\ndoes not match declared type\n" ^
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   381
                            Syntax.string_of_typ_global thy T)) fldTs;
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   382
                     thy)
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   383
                  end)
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   384
       end)
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   385
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   386
  | add_type_def prfx (s, Pending_Type) (ids, thy) =
44030
b63a6bc144cf Pending FDL types may now be associated with Isabelle types as well.
berghofe
parents: 43703
diff changeset
   387
      (ids,
b63a6bc144cf Pending FDL types may now be associated with Isabelle types as well.
berghofe
parents: 43703
diff changeset
   388
       case get_type thy prfx s of
b63a6bc144cf Pending FDL types may now be associated with Isabelle types as well.
berghofe
parents: 43703
diff changeset
   389
         SOME _ => thy
b63a6bc144cf Pending FDL types may now be associated with Isabelle types as well.
berghofe
parents: 43703
diff changeset
   390
       | NONE => Typedecl.typedecl_global
b63a6bc144cf Pending FDL types may now be associated with Isabelle types as well.
berghofe
parents: 43703
diff changeset
   391
           (Binding.name s, [], NoSyn) thy |> snd);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   392
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   393
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   394
fun term_of_expr thy prfx types pfuns =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   395
  let
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   396
    fun tm_of vs (Funct ("->", [e, e'])) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   397
          (HOLogic.mk_imp (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   398
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   399
      | tm_of vs (Funct ("<->", [e, e'])) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   400
          (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   401
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   402
      | tm_of vs (Funct ("or", [e, e'])) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   403
          (HOLogic.mk_disj (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   404
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   405
      | tm_of vs (Funct ("and", [e, e'])) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   406
          (HOLogic.mk_conj (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   407
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   408
      | tm_of vs (Funct ("not", [e])) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   409
          (HOLogic.mk_not (fst (tm_of vs e)), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   410
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   411
      | tm_of vs (Funct ("=", [e, e'])) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   412
          (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   413
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   414
      | tm_of vs (Funct ("<>", [e, e'])) = (HOLogic.mk_not
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   415
          (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e'))), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   416
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   417
      | tm_of vs (Funct ("<", [e, e'])) = (HOLogic.mk_binrel @{const_name less}
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   418
          (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   419
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   420
      | tm_of vs (Funct (">", [e, e'])) = (HOLogic.mk_binrel @{const_name less}
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   421
          (fst (tm_of vs e'), fst (tm_of vs e)), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   422
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   423
      | tm_of vs (Funct ("<=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq}
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   424
          (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   425
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   426
      | tm_of vs (Funct (">=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq}
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   427
          (fst (tm_of vs e'), fst (tm_of vs e)), booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   428
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   429
      | tm_of vs (Funct ("+", [e, e'])) = (HOLogic.mk_binop @{const_name plus}
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   430
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   431
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   432
      | tm_of vs (Funct ("-", [e, e'])) = (HOLogic.mk_binop @{const_name minus}
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   433
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   434
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   435
      | tm_of vs (Funct ("*", [e, e'])) = (HOLogic.mk_binop @{const_name times}
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   436
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   437
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   438
      | tm_of vs (Funct ("/", [e, e'])) = (HOLogic.mk_binop @{const_name divide}
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   439
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   440
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   441
      | tm_of vs (Funct ("div", [e, e'])) = (HOLogic.mk_binop @{const_name sdiv}
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   442
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   443
41635
f938a6022d2e Replaced smod by standard mod operator to reflect actual behaviour
berghofe
parents: 41561
diff changeset
   444
      | tm_of vs (Funct ("mod", [e, e'])) = (HOLogic.mk_binop @{const_name mod}
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   445
          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   446
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   447
      | tm_of vs (Funct ("-", [e])) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   448
          (mk_unop @{const_name uminus} (fst (tm_of vs e)), integerN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   449
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   450
      | tm_of vs (Funct ("**", [e, e'])) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   451
          (Const (@{const_name power}, HOLogic.intT --> HOLogic.natT -->
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   452
             HOLogic.intT) $ fst (tm_of vs e) $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   453
               (@{const nat} $ fst (tm_of vs e')), integerN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   454
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   455
      | tm_of (tab, _) (Ident s) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   456
          (case Symtab.lookup tab s of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   457
             SOME t_ty => t_ty
42499
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
   458
           | NONE => (case lookup_prfx prfx pfuns s of
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
   459
               SOME (SOME ([], resty), t) => (t, resty)
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
   460
             | _ => error ("Undeclared identifier " ^ s)))
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   461
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   462
      | tm_of _ (Number i) = (HOLogic.mk_number HOLogic.intT i, integerN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   463
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   464
      | tm_of vs (Quantifier (s, xs, ty, e)) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   465
          let
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   466
            val (ys, vs') = mk_variables thy prfx xs ty vs;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   467
            val q = (case s of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   468
                "for_all" => HOLogic.mk_all
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   469
              | "for_some" => HOLogic.mk_exists)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   470
          in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   471
            (fold_rev (fn Free (x, T) => fn t => q (x, T, t))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   472
               ys (fst (tm_of vs' e)),
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   473
             booleanN)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   474
          end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   475
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   476
      | tm_of vs (Funct (s, es)) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   477
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   478
          (* record field selection *)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   479
          (case try (unprefix "fld_") s of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   480
             SOME fname => (case es of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   481
               [e] =>
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   482
                 let
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   483
                   val (t, rcdty) = tm_of vs e;
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   484
                   val (rT, cmap) = mk_type' thy prfx rcdty
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   485
                 in case (get_record_info thy rT, lookup types rcdty) of
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   486
                     (SOME {fields, ...}, SOME (Record_Type fldtys)) =>
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   487
                       (case (find_field cmap fname fields,
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   488
                            find_field' fname fldtys) of
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   489
                          (SOME (fname', fT), SOME fldty) =>
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   490
                            (Const (fname', rT --> fT) $ t, fldty)
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   491
                        | _ => error ("Record " ^ rcdty ^
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   492
                            " has no field named " ^ fname))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   493
                   | _ => error (rcdty ^ " is not a record type")
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   494
                 end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   495
             | _ => error ("Function " ^ s ^ " expects one argument"))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   496
           | NONE =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   497
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   498
          (* record field update *)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   499
          (case try (unprefix "upf_") s of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   500
             SOME fname => (case es of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   501
               [e, e'] =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   502
                 let
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   503
                   val (t, rcdty) = tm_of vs e;
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   504
                   val (rT, cmap) = mk_type' thy prfx rcdty;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   505
                   val (u, fldty) = tm_of vs e';
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   506
                   val fT = mk_type thy prfx fldty
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   507
                 in case get_record_info thy rT of
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   508
                     SOME {fields, ...} =>
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   509
                       (case find_field cmap fname fields of
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   510
                          SOME (fname', fU) =>
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   511
                            if fT = fU then
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   512
                              (Const (fname' ^ "_update",
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   513
                                 (fT --> fT) --> rT --> rT) $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   514
                                   Abs ("x", fT, u) $ t,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   515
                               rcdty)
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   516
                            else error ("Type\n" ^
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   517
                              Syntax.string_of_typ_global thy fT ^
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   518
                              "\ndoes not match type\n" ^
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   519
                              Syntax.string_of_typ_global thy fU ^
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   520
                              "\nof field " ^ fname)
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   521
                        | NONE => error ("Record " ^ rcdty ^
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   522
                            " has no field named " ^ fname))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   523
                   | _ => error (rcdty ^ " is not a record type")
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   524
                 end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   525
             | _ => error ("Function " ^ s ^ " expects two arguments"))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   526
           | NONE =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   527
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   528
          (* enumeration type to integer *)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   529
          (case try (unsuffix "__pos") s of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   530
             SOME tyname => (case es of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   531
               [e] => (Const (@{const_name pos},
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   532
                   mk_type thy prfx tyname --> HOLogic.intT) $ fst (tm_of vs e),
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   533
                 integerN)
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   534
             | _ => error ("Function " ^ s ^ " expects one argument"))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   535
           | NONE =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   536
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   537
          (* integer to enumeration type *)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   538
          (case try (unsuffix "__val") s of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   539
             SOME tyname => (case es of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   540
               [e] => (Const (@{const_name val},
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   541
                   HOLogic.intT --> mk_type thy prfx tyname) $ fst (tm_of vs e),
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   542
                 tyname)
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   543
             | _ => error ("Function " ^ s ^ " expects one argument"))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   544
           | NONE =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   545
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   546
          (* successor / predecessor of enumeration type element *)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   547
          if s = "succ" orelse s = "pred" then (case es of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   548
              [e] =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   549
                let
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   550
                  val (t, tyname) = tm_of vs e;
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   551
                  val T = mk_type thy prfx tyname
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   552
                in (Const
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   553
                  (if s = "succ" then @{const_name succ}
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   554
                   else @{const_name pred}, T --> T) $ t, tyname)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   555
                end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   556
            | _ => error ("Function " ^ s ^ " expects one argument"))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   557
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   558
          (* user-defined proof function *)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   559
          else
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   560
            (case lookup_prfx prfx pfuns s of
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   561
               SOME (SOME (_, resty), t) =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   562
                 (list_comb (t, map (fst o tm_of vs) es), resty)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   563
             | _ => error ("Undeclared proof function " ^ s))))))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   564
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   565
      | tm_of vs (Element (e, es)) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   566
          let val (t, ty) = tm_of vs e
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   567
          in case lookup types ty of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   568
              SOME (Array_Type (_, elty)) =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   569
                (t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es), elty)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   570
            | _ => error (ty ^ " is not an array type")
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   571
          end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   572
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   573
      | tm_of vs (Update (e, es, e')) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   574
          let val (t, ty) = tm_of vs e
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   575
          in case lookup types ty of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   576
              SOME (Array_Type (idxtys, elty)) =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   577
                let
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   578
                  val T = foldr1 HOLogic.mk_prodT
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   579
                    (map (mk_type thy prfx) idxtys);
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   580
                  val U = mk_type thy prfx elty;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   581
                  val fT = T --> U
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   582
                in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   583
                  (Const (@{const_name fun_upd}, fT --> T --> U --> fT) $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   584
                     t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es) $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   585
                       fst (tm_of vs e'),
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   586
                   ty)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   587
                end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   588
            | _ => error (ty ^ " is not an array type")
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   589
          end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   590
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   591
      | tm_of vs (Record (s, flds)) =
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   592
          let
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   593
            val (T, cmap) = mk_type' thy prfx s;
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   594
            val {extension = (ext_name, _), fields, ...} =
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   595
              (case get_record_info thy T of
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   596
                 NONE => error (s ^ " is not a record type")
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   597
               | SOME info => info);
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   598
            val flds' = map (apsnd (tm_of vs)) flds;
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   599
            val fnames = fields |> invert_map cmap |>
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   600
              map (Long_Name.base_name o fst);
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   601
            val fnames' = map fst flds;
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   602
            val (fvals, ftys) = split_list (map (fn s' =>
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   603
              case AList.lookup lcase_eq flds' s' of
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   604
                SOME fval_ty => fval_ty
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   605
              | NONE => error ("Field " ^ s' ^ " missing in record " ^ s))
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   606
                  fnames);
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   607
            val _ = (case subtract lcase_eq fnames fnames' of
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   608
                [] => ()
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   609
              | xs => error ("Extra field(s) " ^ commas xs ^
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   610
                  " in record " ^ s));
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   611
            val _ = (case duplicates (op =) fnames' of
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   612
                [] => ()
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   613
              | xs => error ("Duplicate field(s) " ^ commas xs ^
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   614
                  " in record " ^ s))
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   615
          in
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   616
            (list_comb
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   617
               (Const (ext_name,
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   618
                  map (mk_type thy prfx) ftys @ [HOLogic.unitT] ---> T),
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   619
                fvals @ [HOLogic.unit]),
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   620
             s)
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   621
          end
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   622
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   623
      | tm_of vs (Array (s, default, assocs)) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   624
          (case lookup types s of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   625
             SOME (Array_Type (idxtys, elty)) =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   626
               let
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   627
                 val Ts = map (mk_type thy prfx) idxtys;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   628
                 val T = foldr1 HOLogic.mk_prodT Ts;
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   629
                 val U = mk_type thy prfx elty;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   630
                 fun mk_idx' T (e, NONE) = HOLogic.mk_set T [fst (tm_of vs e)]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   631
                   | mk_idx' T (e, SOME e') = Const (@{const_name atLeastAtMost},
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   632
                       T --> T --> HOLogic.mk_setT T) $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   633
                         fst (tm_of vs e) $ fst (tm_of vs e');
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   634
                 fun mk_idx idx =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   635
                   if length Ts <> length idx then
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   636
                     error ("Arity mismatch in construction of array " ^ s)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   637
                   else foldr1 mk_times (map2 mk_idx' Ts idx);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   638
                 fun mk_upd (idxs, e) t =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   639
                   if length idxs = 1 andalso forall (is_none o snd) (hd idxs)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   640
                   then
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   641
                     Const (@{const_name fun_upd}, (T --> U) -->
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   642
                         T --> U --> T --> U) $ t $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   643
                       foldl1 HOLogic.mk_prod
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   644
                         (map (fst o tm_of vs o fst) (hd idxs)) $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   645
                       fst (tm_of vs e)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   646
                   else
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   647
                     Const (@{const_name fun_upds}, (T --> U) -->
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   648
                         HOLogic.mk_setT T --> U --> T --> U) $ t $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   649
                       foldl1 (HOLogic.mk_binop @{const_name sup})
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   650
                         (map mk_idx idxs) $
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   651
                       fst (tm_of vs e)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   652
               in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   653
                 (fold mk_upd assocs
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   654
                    (case default of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   655
                       SOME e => Abs ("x", T, fst (tm_of vs e))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   656
                     | NONE => Const (@{const_name undefined}, T --> U)),
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   657
                  s)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   658
               end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   659
           | _ => error (s ^ " is not an array type"))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   660
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   661
  in tm_of end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   662
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   663
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   664
fun term_of_rule thy prfx types pfuns ids rule =
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   665
  let val tm_of = fst o term_of_expr thy prfx types pfuns ids
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   666
  in case rule of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   667
      Inference_Rule (es, e) => Logic.list_implies
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   668
        (map (HOLogic.mk_Trueprop o tm_of) es, HOLogic.mk_Trueprop (tm_of e))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   669
    | Substitution_Rule (es, e, e') => Logic.list_implies
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   670
        (map (HOLogic.mk_Trueprop o tm_of) es,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   671
         HOLogic.mk_Trueprop (HOLogic.mk_eq (tm_of e, tm_of e')))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   672
  end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   673
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   674
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   675
val builtin = Symtab.make (map (rpair ())
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   676
  ["->", "<->", "or", "and", "not", "=", "<>", "<", ">", "<=", ">=",
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   677
   "+", "-", "*", "/", "div", "mod", "**"]);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   678
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   679
fun complex_expr (Number _) = false
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   680
  | complex_expr (Ident _) = false 
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   681
  | complex_expr (Funct (s, es)) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   682
      not (Symtab.defined builtin s) orelse exists complex_expr es
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   683
  | complex_expr (Quantifier (_, _, _, e)) = complex_expr e
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   684
  | complex_expr _ = true;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   685
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   686
fun complex_rule (Inference_Rule (es, e)) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   687
      complex_expr e orelse exists complex_expr es
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   688
  | complex_rule (Substitution_Rule (es, e, e')) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   689
      complex_expr e orelse complex_expr e' orelse
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   690
      exists complex_expr es;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   691
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   692
val is_pfun =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   693
  Symtab.defined builtin orf
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   694
  can (unprefix "fld_") orf can (unprefix "upf_") orf
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   695
  can (unsuffix "__pos") orf can (unsuffix "__val") orf
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   696
  equal "succ" orf equal "pred";
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   697
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   698
fun fold_opt f = the_default I o Option.map f;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   699
fun fold_pair f g (x, y) = f x #> g y;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   700
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   701
fun fold_expr f g (Funct (s, es)) = f s #> fold (fold_expr f g) es
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   702
  | fold_expr f g (Ident s) = g s
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   703
  | fold_expr f g (Number _) = I
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   704
  | fold_expr f g (Quantifier (_, _, _, e)) = fold_expr f g e
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   705
  | fold_expr f g (Element (e, es)) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   706
      fold_expr f g e #> fold (fold_expr f g) es
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   707
  | fold_expr f g (Update (e, es, e')) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   708
      fold_expr f g e #> fold (fold_expr f g) es #> fold_expr f g e'
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   709
  | fold_expr f g (Record (_, flds)) = fold (fold_expr f g o snd) flds
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   710
  | fold_expr f g (Array (_, default, assocs)) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   711
      fold_opt (fold_expr f g) default #>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   712
      fold (fold_pair
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   713
        (fold (fold (fold_pair
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   714
          (fold_expr f g) (fold_opt (fold_expr f g)))))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   715
        (fold_expr f g)) assocs;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   716
42499
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
   717
fun add_expr_pfuns funs = fold_expr
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
   718
  (fn s => if is_pfun s then I else insert (op =) s)
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
   719
  (fn s => if is_none (lookup funs s) then I else insert (op =) s);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   720
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   721
val add_expr_idents = fold_expr (K I) (insert (op =));
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   722
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   723
fun pfun_type thy prfx (argtys, resty) =
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   724
  map (mk_type thy prfx) argtys ---> mk_type thy prfx resty;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   725
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   726
fun check_pfun_type thy prfx s t optty1 optty2 =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   727
  let
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   728
    val T = fastype_of t;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   729
    fun check ty =
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   730
      let val U = pfun_type thy prfx ty
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   731
      in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   732
        T = U orelse
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   733
        error ("Type\n" ^
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   734
          Syntax.string_of_typ_global thy T ^
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   735
          "\nof function " ^
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   736
          Syntax.string_of_term_global thy t ^
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   737
          " associated with proof function " ^ s ^
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   738
          "\ndoes not match declared type\n" ^
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   739
          Syntax.string_of_typ_global thy U)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   740
      end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   741
  in (Option.map check optty1; Option.map check optty2; ()) end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   742
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   743
fun upd_option x y = if is_some x then x else y;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   744
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   745
fun check_pfuns_types thy prfx funs =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   746
  Symtab.map (fn s => fn (optty, t) =>
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   747
   let val optty' = lookup funs (unprefix_pfun prfx s)
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   748
   in
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   749
     (check_pfun_type thy prfx s t optty optty';
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   750
      (NONE |> upd_option optty |> upd_option optty', t))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   751
   end);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   752
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   753
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   754
(** the VC store **)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   755
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   756
fun pp_vcs msg vcs = Pretty.big_list msg (map (Pretty.str o fst) vcs);
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   757
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   758
fun pp_open_vcs [] = Pretty.str "All verification conditions have been proved."
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   759
  | pp_open_vcs vcs = pp_vcs
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   760
      "The following verification conditions remain to be proved:" vcs;
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   761
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   762
fun partition_vcs vcs = VCtab.fold_rev
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   763
  (fn (name, (trace, SOME thms, ps, cs)) =>
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   764
        apfst (cons (name, (trace, thms, ps, cs)))
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   765
    | (name, (trace, NONE, ps, cs)) =>
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   766
        apsnd (cons (name, (trace, ps, cs))))
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   767
  vcs ([], []);
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   768
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   769
fun insert_break prt = Pretty.blk (0, [Pretty.fbrk, prt]);
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   770
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   771
fun print_open_vcs f vcs =
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   772
  (Pretty.writeln (f (pp_open_vcs (snd (partition_vcs vcs)))); vcs);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   773
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   774
fun set_env ctxt defs types funs ids vcs path prefix thy = VCs.map (fn
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   775
    {pfuns, type_map, env = NONE} =>
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   776
      {pfuns = pfuns,
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   777
       type_map = type_map,
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   778
       env = SOME
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   779
         {ctxt = ctxt, defs = defs, types = types, funs = funs,
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   780
          pfuns = check_pfuns_types thy prefix funs pfuns,
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   781
          ids = ids, proving = false, vcs = print_open_vcs I vcs, path = path,
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   782
          prefix = prefix}}
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   783
  | _ => err_unfinished ()) thy;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   784
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   785
fun mk_pat s = (case Int.fromString s of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   786
    SOME i => [HOLogic.mk_Trueprop (Var (("C", i), HOLogic.boolT))]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   787
  | NONE => error ("Bad conclusion identifier: C" ^ s));
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   788
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   789
fun mk_vc thy prfx types pfuns ids name_concl (tr, proved, ps, cs) =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   790
  let val prop_of =
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   791
    HOLogic.mk_Trueprop o fst o term_of_expr thy prfx types pfuns ids
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   792
  in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   793
    (tr, proved,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   794
     Element.Assumes (map (fn (s', e) =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   795
       ((Binding.name ("H" ^ s'), []), [(prop_of e, [])])) ps),
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   796
     Element.Shows (map (fn (s', e) =>
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   797
       (if name_concl then (Binding.name ("C" ^ s'), [])
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   798
        else Attrib.empty_binding,
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   799
        [(prop_of e, mk_pat s')])) cs))
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   800
  end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   801
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   802
fun fold_vcs f vcs =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   803
  VCtab.fold (fn (_, (_, _, ps, cs)) => fold f ps #> fold f cs) vcs;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   804
42499
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
   805
fun pfuns_of_vcs prfx funs pfuns vcs =
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
   806
  fold_vcs (add_expr_pfuns funs o snd) vcs [] |>
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   807
  filter (is_none o lookup_prfx prfx pfuns);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   808
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   809
fun declare_missing_pfuns thy prfx funs pfuns vcs (tab, ctxt) =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   810
  let
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   811
    val (fs, (tys, Ts)) =
42499
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
   812
      pfuns_of_vcs prfx funs pfuns vcs |>
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   813
      map_filter (fn s => lookup funs s |>
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   814
        Option.map (fn ty => (s, (SOME ty, pfun_type thy prfx ty)))) |>
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   815
      split_list ||> split_list;
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42499
diff changeset
   816
    val (fs', ctxt') = fold_map Name.variant fs ctxt
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   817
  in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   818
    (fold Symtab.update_new (fs ~~ (tys ~~ map Free (fs' ~~ Ts))) pfuns,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   819
     Element.Fixes (map2 (fn s => fn T =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   820
       (Binding.name s, SOME T, NoSyn)) fs' Ts),
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   821
     (tab, ctxt'))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   822
  end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   823
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   824
fun map_pfuns f {pfuns, type_map, env} =
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   825
  {pfuns = f pfuns, type_map = type_map, env = env}
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   826
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   827
fun map_pfuns_env f {pfuns, type_map,
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   828
      env = SOME {ctxt, defs, types, funs, pfuns = pfuns_env,
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   829
        ids, proving, vcs, path, prefix}} =
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   830
  if proving then err_unfinished ()
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   831
  else
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   832
    {pfuns = pfuns, type_map = type_map,
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   833
     env = SOME {ctxt = ctxt, defs = defs, types = types, funs = funs,
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   834
       pfuns = f pfuns_env, ids = ids, proving = false, vcs = vcs,
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   835
       path = path, prefix = prefix}};
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   836
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   837
fun add_proof_fun prep (s, (optty, raw_t)) thy =
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   838
  VCs.map (fn data as {env, ...} =>
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   839
    let
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   840
      val (optty', prfx, map_pf) = (case env of
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   841
          SOME {funs, prefix, ...} =>
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   842
            (lookup funs (unprefix_pfun prefix s),
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   843
             prefix, map_pfuns_env)
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   844
        | NONE => (NONE, "", map_pfuns));
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   845
      val optty'' = NONE |> upd_option optty |> upd_option optty';
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   846
      val t = prep (Option.map (pfun_type thy prfx) optty'') raw_t;
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   847
      val _ = (case fold_aterms (fn u =>
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   848
          if is_Var u orelse is_Free u then insert (op =) u else I) t [] of
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   849
          [] => ()
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   850
        | ts => error ("Term\n" ^ Syntax.string_of_term_global thy t ^
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   851
            "\nto be associated with proof function " ^ s ^
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   852
            " contains free variable(s) " ^
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   853
            commas (map (Syntax.string_of_term_global thy) ts)));
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   854
    in
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   855
      (check_pfun_type thy prfx s t optty optty';
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   856
       if is_some optty'' orelse is_none env then
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   857
         map_pf (Symtab.update_new (s, (optty'', t))) data
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   858
           handle Symtab.DUP _ => error ("Proof function " ^ s ^
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   859
             " already associated with function")
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   860
       else error ("Undeclared proof function " ^ s))
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   861
    end) thy;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   862
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   863
fun check_mapping _ _ [] _ = ()
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   864
  | check_mapping err s cmap cs =
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58956
diff changeset
   865
      (case duplicates (op = o apply2 fst) cmap of
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58956
diff changeset
   866
         [] => (case duplicates (op = o apply2 snd) cmap of
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   867
             [] => (case subtract (op = o apsnd snd) cs cmap of
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   868
                 [] => (case subtract (op = o apfst snd) cmap cs of
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   869
                     [] => ()
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   870
                   | zs => err ("has extra " ^ s ^ "(s) " ^ commas zs))
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   871
               | zs => err ("does not have " ^ s ^ "(s) " ^
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   872
                   commas (map snd zs)))
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   873
           | zs => error ("Several SPARK names are mapped to " ^
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   874
               commas (map snd zs)))
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   875
       | zs => error ("The SPARK names " ^ commas (map fst zs) ^
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   876
           " are mapped to more than one name"));
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   877
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   878
fun add_type (s, (T as Type (tyname, Ts), cmap)) thy =
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   879
      let val cmap' = map (apsnd (Sign.intern_const thy)) cmap
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   880
      in
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   881
        thy |>
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   882
        VCs.map (fn
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   883
            {env = SOME _, ...} => err_unfinished ()
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   884
          | {pfuns, type_map, env} =>
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   885
              {pfuns = pfuns,
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   886
               type_map = Symtab.update_new (s, (T, cmap')) type_map,
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   887
               env = env}
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   888
                handle Symtab.DUP _ => error ("SPARK type " ^ s ^
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   889
                  " already associated with type")) |>
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   890
        (fn thy' =>
58130
5e9170812356 ported to use new-style datatypes
blanchet
parents: 58112
diff changeset
   891
           case BNF_LFP_Compat.get_constrs thy' tyname of
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   892
             NONE => (case get_record_info thy' T of
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   893
               NONE => thy'
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   894
             | SOME {fields, ...} =>
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   895
                 (check_mapping (assoc_ty_err thy' T s) "field"
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   896
                    cmap' (map fst fields);
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   897
                  thy'))
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   898
           | SOME cs =>
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   899
               if null Ts then
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   900
                 (map
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   901
                    (fn (_, Type (_, [])) => ()
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   902
                      | (cname, _) => assoc_ty_err thy' T s
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   903
                          ("has illegal constructor " ^
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   904
                           Long_Name.base_name cname)) cs;
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   905
                  check_mapping (assoc_ty_err thy' T s) "constructor"
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   906
                    cmap' (map fst cs);
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   907
                  add_enum_type s tyname thy')
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   908
               else assoc_ty_err thy' T s "is illegal")
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   909
      end
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 46724
diff changeset
   910
  | add_type (s, (T, _)) thy = assoc_ty_err thy T s "is illegal";
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   911
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   912
val is_closed = is_none o #env o VCs.get;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   913
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   914
fun lookup_vc thy name_concl name =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   915
  (case VCs.get thy of
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   916
    {env = SOME {vcs, types, funs, pfuns, ids, ctxt, prefix, ...}, ...} =>
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   917
      (case VCtab.lookup vcs name of
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   918
         SOME vc =>
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   919
           let val (pfuns', ctxt', ids') =
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   920
             declare_missing_pfuns thy prefix funs pfuns vcs ids
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   921
           in
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   922
             SOME (ctxt @ [ctxt'],
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   923
               mk_vc thy prefix types pfuns' ids' name_concl vc)
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   924
           end
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   925
       | NONE => NONE)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   926
  | _ => NONE);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   927
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   928
fun get_vcs thy name_concl = (case VCs.get thy of
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   929
    {env = SOME {vcs, types, funs, pfuns, ids, ctxt, defs, prefix, ...}, ...} =>
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   930
      let val (pfuns', ctxt', ids') =
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   931
        declare_missing_pfuns thy prefix funs pfuns vcs ids
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   932
      in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   933
        (ctxt @ [ctxt'], defs,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   934
         VCtab.dest vcs |>
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   935
         map (apsnd (mk_vc thy prefix types pfuns' ids' name_concl)))
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   936
      end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   937
  | _ => ([], [], []));
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   938
41896
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41878
diff changeset
   939
fun mark_proved name thms = VCs.map (fn
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   940
    {pfuns, type_map,
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   941
     env = SOME {ctxt, defs, types, funs, pfuns = pfuns_env,
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   942
        ids, vcs, path, prefix, ...}} =>
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   943
      {pfuns = pfuns,
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
   944
       type_map = type_map,
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   945
       env = SOME {ctxt = ctxt, defs = defs,
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   946
         types = types, funs = funs, pfuns = pfuns_env,
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
   947
         ids = ids,
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   948
         proving = true,
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   949
         vcs = print_open_vcs insert_break (VCtab.map_entry name
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   950
           (fn (trace, _, ps, cs) => (trace, SOME thms, ps, cs)) vcs),
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   951
         path = path,
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
   952
         prefix = prefix}}
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   953
  | x => x);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   954
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   955
fun close incomplete thy =
59810
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   956
  thy |> VCs.map (fn {pfuns, type_map, env} =>
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   957
    (case env of
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   958
      NONE => error "No SPARK environment is currently open"
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   959
    | SOME {vcs, path, ...} =>
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   960
        let
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   961
          val (proved, unproved) = partition_vcs vcs;
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   962
          val _ = Thm.join_proofs (maps (#2 o snd) proved);
59810
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   963
          val (proved', proved'') =
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   964
            List.partition (fn (_, (_, thms, _, _)) =>
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   965
              exists (#oracle o Thm.peek_status) thms) proved;
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   966
          val _ =
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   967
            if null unproved then ()
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   968
            else (if incomplete then warning else error) (Pretty.string_of (pp_open_vcs unproved));
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   969
          val _ =
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   970
            if null proved' then ()
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   971
            else warning (Pretty.string_of (pp_vcs
48167
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   972
             "The following VCs are not marked as proved \
da1a1eae93fa Various improvements
berghofe
parents: 47880
diff changeset
   973
             \because their proofs contain oracles:" proved'));
59810
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   974
          val prv_path = Path.ext "prv" path;
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   975
          val _ =
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   976
            if File.exists prv_path orelse Options.default_bool @{system_option spark_prv} then
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   977
              File.write prv_path
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   978
               (implode (map (fn (s, _) => snd (strip_number s) ^
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   979
                  " -- proved by " ^ Distribution.version ^ "\n") proved''))
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   980
            else ();
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   981
        in {pfuns = pfuns, type_map = type_map, env = NONE} end))
e749a0f2f401 HOL-SPARK .prv files are subject to system option spark_prv;
wenzelm
parents: 59498
diff changeset
   982
  |> Sign.parent_path;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   983
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   984
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   985
(** set up verification conditions **)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   986
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   987
fun partition_opt f =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   988
  let
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   989
    fun part ys zs [] = (rev ys, rev zs)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   990
      | part ys zs (x :: xs) = (case f x of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   991
            SOME y => part (y :: ys) zs xs
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   992
          | NONE => part ys (x :: zs) xs)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   993
  in part [] [] end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   994
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   995
fun dest_def (id, (Substitution_Rule ([], Ident s, rhs))) = SOME (id, (s, rhs))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   996
  | dest_def _ = NONE;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   997
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   998
fun mk_rulename (s, i) = Binding.name (s ^ string_of_int i);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   999
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1000
fun add_const prfx (s, ty) ((tab, ctxt), thy) =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1001
  let
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1002
    val T = mk_type thy prfx ty;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1003
    val b = Binding.name s;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1004
    val c = Const (Sign.full_name thy b, T)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1005
  in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1006
    (c,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1007
     ((Symtab.update (s, (c, ty)) tab, Name.declare s ctxt),
56239
17df7145a871 tuned signature;
wenzelm
parents: 51717
diff changeset
  1008
      Sign.add_consts [(b, T, NoSyn)] thy))
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1009
  end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1010
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1011
fun add_def prfx types pfuns consts (id, (s, e)) (ids as (tab, ctxt), thy) =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1012
  (case lookup consts s of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1013
     SOME ty =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1014
       let
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1015
         val (t, ty') = term_of_expr thy prfx types pfuns ids e;
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1016
         val T = mk_type thy prfx ty;
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1017
         val T' = mk_type thy prfx ty';
41878
0778ade00b06 - Made sure that sort_defs is aware of constants introduced by add_type_def
berghofe
parents: 41635
diff changeset
  1018
         val _ = T = T' orelse
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1019
           error ("Declared type " ^ ty ^ " of " ^ s ^
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1020
             "\ndoes not match type " ^ ty' ^ " in definition");
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1021
         val id' = mk_rulename id;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1022
         val lthy = Named_Target.theory_init thy;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1023
         val ((t', (_, th)), lthy') = Specification.definition
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1024
           (NONE, ((id', []), HOLogic.mk_Trueprop (HOLogic.mk_eq
41878
0778ade00b06 - Made sure that sort_defs is aware of constants introduced by add_type_def
berghofe
parents: 41635
diff changeset
  1025
             (Free (s, T), t)))) lthy;
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42356
diff changeset
  1026
         val phi = Proof_Context.export_morphism lthy' lthy
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1027
       in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1028
         ((id', Morphism.thm phi th),
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1029
          ((Symtab.update (s, (Morphism.term phi t', ty)) tab,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1030
            Name.declare s ctxt),
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1031
           Local_Theory.exit_global lthy'))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1032
       end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1033
   | NONE => error ("Undeclared constant " ^ s));
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1034
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1035
fun add_var prfx (s, ty) (ids, thy) =
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1036
  let val ([Free p], ids') = mk_variables thy prfx [s] ty ids
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1037
  in (p, (ids', thy)) end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1038
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1039
fun add_init_vars prfx vcs (ids_thy as ((tab, _), _)) =
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1040
  fold_map (add_var prfx)
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1041
    (map_filter
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1042
       (fn s => case try (unsuffix "~") s of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1043
          SOME s' => (case Symtab.lookup tab s' of
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1044
            SOME (_, ty) => SOME (s, ty)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1045
          | NONE => error ("Undeclared identifier " ^ s'))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1046
        | NONE => NONE)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1047
       (fold_vcs (add_expr_idents o snd) vcs []))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1048
    ids_thy;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1049
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1050
fun is_trivial_vc ([], [(_, Ident "true")]) = true
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1051
  | is_trivial_vc _ = false;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1052
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1053
fun rulenames rules = commas
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1054
  (map (fn ((s, i), _) => s ^ "(" ^ string_of_int i ^ ")") rules);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1055
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1056
(* sort definitions according to their dependency *)
42499
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
  1057
fun sort_defs _ _ _ _ [] sdefs = rev sdefs
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
  1058
  | sort_defs prfx funs pfuns consts defs sdefs =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1059
      (case find_first (fn (_, (_, e)) =>
42499
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
  1060
         forall (is_some o lookup_prfx prfx pfuns)
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
  1061
           (add_expr_pfuns funs e []) andalso
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1062
         forall (fn id =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1063
           member (fn (s, (_, (s', _))) => s = s') sdefs id orelse
41878
0778ade00b06 - Made sure that sort_defs is aware of constants introduced by add_type_def
berghofe
parents: 41635
diff changeset
  1064
           consts id)
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1065
             (add_expr_idents e [])) defs of
42499
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
  1066
         SOME d => sort_defs prfx funs pfuns consts
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1067
           (remove (op =) d defs) (d :: sdefs)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1068
       | NONE => error ("Bad definitions: " ^ rulenames defs));
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1069
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
  1070
fun set_vcs ({types, vars, consts, funs} : decls)
48453
2421ff8c57a5 set_vcs now derives prefix from fully qualified procedure / function name
berghofe
parents: 48167
diff changeset
  1071
      (rules, _) ((_, ident), vcs) path opt_prfx thy =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1072
  let
48453
2421ff8c57a5 set_vcs now derives prefix from fully qualified procedure / function name
berghofe
parents: 48167
diff changeset
  1073
    val prfx' =
2421ff8c57a5 set_vcs now derives prefix from fully qualified procedure / function name
berghofe
parents: 48167
diff changeset
  1074
      if opt_prfx = "" then
2421ff8c57a5 set_vcs now derives prefix from fully qualified procedure / function name
berghofe
parents: 48167
diff changeset
  1075
        space_implode "__" (Long_Name.explode (Long_Name.qualifier ident))
2421ff8c57a5 set_vcs now derives prefix from fully qualified procedure / function name
berghofe
parents: 48167
diff changeset
  1076
      else opt_prfx;
2421ff8c57a5 set_vcs now derives prefix from fully qualified procedure / function name
berghofe
parents: 48167
diff changeset
  1077
    val prfx = to_lower prfx';
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1078
    val {pfuns, ...} = VCs.get thy;
41878
0778ade00b06 - Made sure that sort_defs is aware of constants introduced by add_type_def
berghofe
parents: 41635
diff changeset
  1079
    val (defs, rules') = partition_opt dest_def rules;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1080
    val consts' =
41878
0778ade00b06 - Made sure that sort_defs is aware of constants introduced by add_type_def
berghofe
parents: 41635
diff changeset
  1081
      subtract (fn ((_, (s, _)), (s', _)) => s = s') defs (items consts);
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1082
    (* ignore all complex rules in rls files *)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1083
    val (rules'', other_rules) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1084
      List.partition (complex_rule o snd) rules';
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1085
    val _ = if null rules'' then ()
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1086
      else warning ("Ignoring rules: " ^ rulenames rules'');
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1087
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1088
    val vcs' = VCtab.make (maps (fn (tr, vcs) =>
41896
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41878
diff changeset
  1089
      map (fn (s, (ps, cs)) => (s, (tr, NONE, ps, cs)))
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1090
        (filter_out (is_trivial_vc o snd) vcs)) vcs);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1091
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1092
    val _ = (case filter_out (is_some o lookup funs)
42499
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
  1093
        (pfuns_of_vcs prfx funs pfuns vcs') of
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1094
        [] => ()
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1095
      | fs => error ("Undeclared proof function(s) " ^ commas fs));
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1096
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1097
    val (((defs', vars''), ivars), (ids, thy')) =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1098
      ((Symtab.empty |>
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45701
diff changeset
  1099
        Symtab.update ("false", (@{term False}, booleanN)) |>
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45701
diff changeset
  1100
        Symtab.update ("true", (@{term True}, booleanN)),
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 41896
diff changeset
  1101
        Name.context),
46724
5759ecd5c905 Use long prefix rather than short package name to disambiguate constant names
berghofe
parents: 46567
diff changeset
  1102
       thy |> Sign.add_path
48453
2421ff8c57a5 set_vcs now derives prefix from fully qualified procedure / function name
berghofe
parents: 48167
diff changeset
  1103
         ((if prfx' = "" then "" else prfx' ^ "__") ^ Long_Name.base_name ident)) |>
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1104
      fold (add_type_def prfx) (items types) |>
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1105
      fold (snd oo add_const prfx) consts' |> (fn ids_thy as ((tab, _), _) =>
41878
0778ade00b06 - Made sure that sort_defs is aware of constants introduced by add_type_def
berghofe
parents: 41635
diff changeset
  1106
        ids_thy |>
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1107
        fold_map (add_def prfx types pfuns consts)
42499
adfa6ad43ab6 Properly treat proof functions with no arguments.
berghofe
parents: 42440
diff changeset
  1108
          (sort_defs prfx funs pfuns (Symtab.defined tab) defs []) ||>>
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1109
        fold_map (add_var prfx) (items vars) ||>>
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1110
        add_init_vars prfx vcs');
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1111
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1112
    val ctxt =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1113
      [Element.Fixes (map (fn (s, T) =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1114
         (Binding.name s, SOME T, NoSyn)) (vars'' @ ivars)),
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1115
       Element.Assumes (map (fn (id, rl) =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1116
         ((mk_rulename id, []),
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
  1117
          [(term_of_rule thy' prfx types pfuns ids rl, [])]))
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1118
           other_rules),
42440
5e7a7343ab11 discontinuend obsolete Thm.definitionK, which was hardly ever well-defined;
wenzelm
parents: 42416
diff changeset
  1119
       Element.Notes ("", [((Binding.name "defns", []), map (rpair [] o single o snd) defs')])]
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1120
          
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1121
  in
46567
8421b6cf2a33 Fixed bugs:
berghofe
parents: 45906
diff changeset
  1122
    set_env ctxt defs' types funs ids vcs' path prfx thy'
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1123
  end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1124
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
  1125
end;