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