berghofe@41809: (* Title: HOL/SPARK/Tools/spark_vcs.ML berghofe@41809: Author: Stefan Berghofer berghofe@41809: Copyright: secunet Security Networks AG berghofe@41809: berghofe@41809: Store for verification conditions generated by SPARK/Ada. berghofe@41809: *) berghofe@41809: berghofe@41809: signature SPARK_VCS = berghofe@41809: sig berghofe@43227: val set_vcs: Fdl_Parser.decls -> Fdl_Parser.rules -> berghofe@43267: (string * string) * Fdl_Parser.vcs -> Path.T -> string -> theory -> theory berghofe@41809: val add_proof_fun: (typ option -> 'a -> term) -> berghofe@41809: string * ((string list * string) option * 'a) -> berghofe@41809: theory -> theory berghofe@47596: val add_type: string * (typ * (string * string) list) -> theory -> theory berghofe@49182: val lookup_vc: theory -> bool -> string -> (Element.context_i list * berghofe@42767: (string * thm list option * Element.context_i * Element.statement_i)) option berghofe@49182: val get_vcs: theory -> bool -> berghofe@42767: Element.context_i list * (binding * thm) list * (string * berghofe@42767: (string * thm list option * Element.context_i * Element.statement_i)) list berghofe@42767: val mark_proved: string -> thm list -> theory -> theory berghofe@49182: val close: bool -> theory -> theory berghofe@41809: val is_closed: theory -> bool berghofe@41809: end; berghofe@41809: berghofe@41809: structure SPARK_VCs: SPARK_VCS = berghofe@41809: struct berghofe@41809: berghofe@41809: open Fdl_Parser; berghofe@41809: berghofe@41809: berghofe@43227: (** theory data **) berghofe@43227: berghofe@43227: fun err_unfinished () = error "An unfinished SPARK environment is still open." berghofe@43227: berghofe@43227: val strip_number = pairself implode o take_suffix Fdl_Lexer.is_digit o raw_explode; berghofe@43227: berghofe@43227: val name_ord = prod_ord string_ord (option_ord int_ord) o berghofe@43227: pairself (strip_number ##> Int.fromString); berghofe@43227: berghofe@43227: structure VCtab = Table(type key = string val ord = name_ord); berghofe@43227: berghofe@43227: structure VCs = Theory_Data berghofe@43227: ( berghofe@43227: type T = berghofe@43227: {pfuns: ((string list * string) option * term) Symtab.table, berghofe@47596: type_map: (typ * (string * string) list) Symtab.table, berghofe@43227: env: berghofe@43227: {ctxt: Element.context_i list, berghofe@43227: defs: (binding * thm) list, berghofe@43227: types: fdl_type tab, berghofe@43227: funs: (string list * string) tab, berghofe@47438: pfuns: ((string list * string) option * term) Symtab.table, berghofe@43227: ids: (term * string) Symtab.table * Name.context, berghofe@43227: proving: bool, berghofe@43227: vcs: (string * thm list option * berghofe@43227: (string * expr) list * (string * expr) list) VCtab.table, berghofe@43267: path: Path.T, berghofe@43267: prefix: string} option} berghofe@43227: val empty : T = {pfuns = Symtab.empty, type_map = Symtab.empty, env = NONE} berghofe@43227: val extend = I berghofe@43227: fun merge ({pfuns = pfuns1, type_map = type_map1, env = NONE}, berghofe@43227: {pfuns = pfuns2, type_map = type_map2, env = NONE}) = berghofe@43227: {pfuns = Symtab.merge (eq_pair (op =) (op aconv)) (pfuns1, pfuns2), berghofe@43227: type_map = Symtab.merge (op =) (type_map1, type_map2), berghofe@43227: env = NONE} berghofe@43227: | merge _ = err_unfinished () berghofe@43227: ) berghofe@43227: berghofe@43227: berghofe@41809: (** utilities **) berghofe@41809: berghofe@43227: val to_lower = raw_explode #> map Symbol.to_ascii_lower #> implode; berghofe@43227: berghofe@43227: val lcase_eq = (op =) o pairself (to_lower o Long_Name.base_name); berghofe@43227: berghofe@43267: fun lookup_prfx "" tab s = Symtab.lookup tab s berghofe@43267: | lookup_prfx prfx tab s = (case Symtab.lookup tab s of berghofe@43267: NONE => Symtab.lookup tab (prfx ^ "__" ^ s) berghofe@43267: | x => x); berghofe@43267: berghofe@43267: fun strip_prfx s = berghofe@43267: let berghofe@43267: fun strip ys [] = ("", implode ys) berghofe@43267: | strip ys ("_" :: "_" :: xs) = (implode (rev xs), implode ys) berghofe@43267: | strip ys (x :: xs) = strip (x :: ys) xs berghofe@43267: in strip [] (rev (raw_explode s)) end; berghofe@43267: berghofe@47596: fun unprefix_pfun "" s = s berghofe@47596: | unprefix_pfun prfx s = berghofe@47596: let val (prfx', s') = strip_prfx s berghofe@47596: in if prfx = prfx' then s' else s end; berghofe@47596: berghofe@41809: fun mk_unop s t = berghofe@41809: let val T = fastype_of t berghofe@41809: in Const (s, T --> T) $ t end; berghofe@41809: berghofe@41809: fun mk_times (t, u) = berghofe@41809: let berghofe@41809: val setT = fastype_of t; berghofe@41809: val T = HOLogic.dest_setT setT; berghofe@41809: val U = HOLogic.dest_setT (fastype_of u) berghofe@41809: in berghofe@41809: Const (@{const_name Sigma}, setT --> (T --> HOLogic.mk_setT U) --> berghofe@41809: HOLogic.mk_setT (HOLogic.mk_prodT (T, U))) $ t $ Abs ("", T, u) berghofe@41809: end; berghofe@41809: berghofe@43267: fun get_type thy prfx ty = berghofe@43227: let val {type_map, ...} = VCs.get thy berghofe@43267: in lookup_prfx prfx type_map ty end; berghofe@43227: berghofe@47596: fun mk_type' _ _ "integer" = (HOLogic.intT, []) berghofe@47596: | mk_type' _ _ "boolean" = (HOLogic.boolT, []) berghofe@47596: | mk_type' thy prfx ty = berghofe@43267: (case get_type thy prfx ty of berghofe@43227: NONE => berghofe@47596: (Syntax.check_typ (Proof_Context.init_global thy) berghofe@47596: (Type (Sign.full_name thy (Binding.name ty), [])), berghofe@47596: []) berghofe@47596: | SOME p => p); berghofe@47596: berghofe@47596: fun mk_type thy prfx ty = fst (mk_type' thy prfx ty); berghofe@41809: berghofe@41809: val booleanN = "boolean"; berghofe@41809: val integerN = "integer"; berghofe@41809: berghofe@41809: fun define_overloaded (def_name, eq) lthy = berghofe@41809: let berghofe@41809: val ((c, _), rhs) = eq |> Syntax.check_term lthy |> berghofe@41809: Logic.dest_equals |>> dest_Free; berghofe@41809: val ((_, (_, thm)), lthy') = Local_Theory.define berghofe@41809: ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs)) lthy wenzelm@43232: val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy'); wenzelm@43232: val thm' = singleton (Proof_Context.export lthy' ctxt_thy) thm berghofe@41809: in (thm', lthy') end; berghofe@41809: berghofe@41809: fun strip_underscores s = berghofe@41809: strip_underscores (unsuffix "_" s) handle Fail _ => s; berghofe@41809: berghofe@41809: fun strip_tilde s = berghofe@41809: unsuffix "~" s ^ "_init" handle Fail _ => s; berghofe@41809: berghofe@41809: val mangle_name = strip_underscores #> strip_tilde; berghofe@41809: berghofe@43267: fun mk_variables thy prfx xs ty (tab, ctxt) = berghofe@41809: let berghofe@43267: val T = mk_type thy prfx ty; wenzelm@44208: val (ys, ctxt') = fold_map Name.variant (map mangle_name xs) ctxt; berghofe@41809: val zs = map (Free o rpair T) ys; berghofe@41809: in (zs, (fold (Symtab.update o apsnd (rpair ty)) (xs ~~ zs) tab, ctxt')) end; berghofe@41809: berghofe@43227: fun get_record_info thy T = (case Record.dest_recTs T of berghofe@43227: [(tyname, [@{typ unit}])] => berghofe@43227: Record.get_info thy (Long_Name.qualifier tyname) berghofe@43227: | _ => NONE); berghofe@43227: berghofe@47596: fun find_field [] fname fields = berghofe@47596: find_first (curry lcase_eq fname o fst) fields berghofe@47596: | find_field cmap fname fields = berghofe@47596: (case AList.lookup (op =) cmap fname of berghofe@47596: NONE => NONE berghofe@47596: | SOME fname' => SOME (fname', the (AList.lookup (op =) fields fname'))); berghofe@43227: berghofe@43227: fun find_field' fname = get_first (fn (flds, fldty) => berghofe@43227: if member (op =) flds fname then SOME fldty else NONE); berghofe@43227: berghofe@43227: fun assoc_ty_err thy T s msg = berghofe@43227: error ("Type " ^ Syntax.string_of_typ_global thy T ^ berghofe@43227: " associated with SPARK type " ^ s ^ "\n" ^ msg); berghofe@43227: berghofe@41809: berghofe@41809: (** generate properties of enumeration types **) berghofe@41809: berghofe@43227: fun add_enum_type tyname tyname' thy = berghofe@41809: let wenzelm@46777: val {case_name, ...} = the (Datatype.get_info thy tyname'); wenzelm@46777: val cs = map Const (the (Datatype.get_constrs thy tyname')); berghofe@43227: val k = length cs; berghofe@41809: val T = Type (tyname', []); berghofe@41809: val p = Const (@{const_name pos}, T --> HOLogic.intT); berghofe@41809: val v = Const (@{const_name val}, HOLogic.intT --> T); berghofe@41809: val card = Const (@{const_name card}, berghofe@41809: HOLogic.mk_setT T --> HOLogic.natT) $ HOLogic.mk_UNIV T; berghofe@41809: berghofe@41809: fun mk_binrel_def s f = Logic.mk_equals berghofe@41809: (Const (s, T --> T --> HOLogic.boolT), berghofe@41809: Abs ("x", T, Abs ("y", T, berghofe@41809: Const (s, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ berghofe@41809: (f $ Bound 1) $ (f $ Bound 0)))); berghofe@41809: berghofe@41809: val (((def1, def2), def3), lthy) = thy |> berghofe@41809: berghofe@43273: Class.instantiation ([tyname'], [], @{sort spark_enum}) |> berghofe@41809: berghofe@41809: define_overloaded ("pos_" ^ tyname ^ "_def", Logic.mk_equals berghofe@41809: (p, berghofe@41809: list_comb (Const (case_name, replicate k HOLogic.intT @ berghofe@41809: [T] ---> HOLogic.intT), berghofe@41809: map (HOLogic.mk_number HOLogic.intT) (0 upto k - 1)))) ||>> berghofe@41809: berghofe@41809: define_overloaded ("less_eq_" ^ tyname ^ "_def", berghofe@41809: mk_binrel_def @{const_name less_eq} p) ||>> berghofe@41809: define_overloaded ("less_" ^ tyname ^ "_def", berghofe@41809: mk_binrel_def @{const_name less} p); berghofe@41809: berghofe@41809: val UNIV_eq = Goal.prove lthy [] [] berghofe@41809: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@41809: (HOLogic.mk_UNIV T, HOLogic.mk_set T cs))) berghofe@41809: (fn _ => berghofe@41809: rtac @{thm subset_antisym} 1 THEN berghofe@41809: rtac @{thm subsetI} 1 THEN wenzelm@46777: Datatype_Aux.exh_tac (K (#exhaust (Datatype.the_info wenzelm@43232: (Proof_Context.theory_of lthy) tyname'))) 1 THEN berghofe@41809: ALLGOALS (asm_full_simp_tac (simpset_of lthy))); berghofe@41809: berghofe@41809: val finite_UNIV = Goal.prove lthy [] [] berghofe@41809: (HOLogic.mk_Trueprop (Const (@{const_name finite}, berghofe@41809: HOLogic.mk_setT T --> HOLogic.boolT) $ HOLogic.mk_UNIV T)) berghofe@41809: (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1); berghofe@41809: berghofe@41809: val card_UNIV = Goal.prove lthy [] [] berghofe@41809: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@41809: (card, HOLogic.mk_number HOLogic.natT k))) berghofe@41809: (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1); berghofe@41809: berghofe@41809: val range_pos = Goal.prove lthy [] [] berghofe@41809: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@41809: (Const (@{const_name image}, (T --> HOLogic.intT) --> berghofe@41809: HOLogic.mk_setT T --> HOLogic.mk_setT HOLogic.intT) $ berghofe@41809: p $ HOLogic.mk_UNIV T, berghofe@41809: Const (@{const_name atLeastLessThan}, HOLogic.intT --> berghofe@41809: HOLogic.intT --> HOLogic.mk_setT HOLogic.intT) $ berghofe@41809: HOLogic.mk_number HOLogic.intT 0 $ berghofe@41809: (@{term int} $ card)))) berghofe@41809: (fn _ => berghofe@41809: simp_tac (simpset_of lthy addsimps [card_UNIV]) 1 THEN berghofe@41809: simp_tac (simpset_of lthy addsimps [UNIV_eq, def1]) 1 THEN berghofe@41809: rtac @{thm subset_antisym} 1 THEN berghofe@41809: simp_tac (simpset_of lthy) 1 THEN berghofe@41809: rtac @{thm subsetI} 1 THEN berghofe@41809: asm_full_simp_tac (simpset_of lthy addsimps @{thms interval_expand} berghofe@41809: delsimps @{thms atLeastLessThan_iff}) 1); berghofe@41809: berghofe@41809: val lthy' = berghofe@41809: Class.prove_instantiation_instance (fn _ => berghofe@41809: Class.intro_classes_tac [] THEN berghofe@41809: rtac finite_UNIV 1 THEN berghofe@41809: rtac range_pos 1 THEN berghofe@41809: simp_tac (HOL_basic_ss addsimps [def3]) 1 THEN berghofe@41809: simp_tac (HOL_basic_ss addsimps [def2]) 1) lthy; berghofe@41809: berghofe@41809: val (pos_eqs, val_eqs) = split_list (map_index (fn (i, c) => berghofe@41809: let berghofe@41809: val n = HOLogic.mk_number HOLogic.intT i; berghofe@41809: val th = Goal.prove lthy' [] [] berghofe@41809: (HOLogic.mk_Trueprop (HOLogic.mk_eq (p $ c, n))) berghofe@41809: (fn _ => simp_tac (simpset_of lthy' addsimps [def1]) 1); berghofe@41809: val th' = Goal.prove lthy' [] [] berghofe@41809: (HOLogic.mk_Trueprop (HOLogic.mk_eq (v $ n, c))) berghofe@41809: (fn _ => berghofe@41809: rtac (@{thm inj_pos} RS @{thm injD}) 1 THEN berghofe@41809: simp_tac (simpset_of lthy' addsimps berghofe@41809: [@{thm pos_val}, range_pos, card_UNIV, th]) 1) berghofe@41809: in (th, th') end) cs); berghofe@41809: berghofe@41809: val first_el = Goal.prove lthy' [] [] berghofe@41809: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@41809: (Const (@{const_name first_el}, T), hd cs))) berghofe@41809: (fn _ => simp_tac (simpset_of lthy' addsimps berghofe@41809: [@{thm first_el_def}, hd val_eqs]) 1); berghofe@41809: berghofe@41809: val last_el = Goal.prove lthy' [] [] berghofe@41809: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@41809: (Const (@{const_name last_el}, T), List.last cs))) berghofe@41809: (fn _ => simp_tac (simpset_of lthy' addsimps berghofe@41809: [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1); berghofe@41809: in berghofe@43227: lthy' |> berghofe@43227: Local_Theory.note wenzelm@46463: ((Binding.name (tyname ^ "_card"), @{attributes [simp]}), [card_UNIV]) ||>> berghofe@43227: Local_Theory.note wenzelm@46463: ((Binding.name (tyname ^ "_pos"), @{attributes [simp]}), pos_eqs) ||>> berghofe@43227: Local_Theory.note wenzelm@46463: ((Binding.name (tyname ^ "_val"), @{attributes [simp]}), val_eqs) ||>> berghofe@43227: Local_Theory.note wenzelm@46463: ((Binding.name (tyname ^ "_first_el"), @{attributes [simp]}), [first_el]) ||>> berghofe@43227: Local_Theory.note wenzelm@46463: ((Binding.name (tyname ^ "_last_el"), @{attributes [simp]}), [last_el]) |> snd |> berghofe@43227: Local_Theory.exit_global berghofe@41809: end; berghofe@41809: berghofe@41809: berghofe@43267: fun check_no_assoc thy prfx s = case get_type thy prfx s of berghofe@43227: NONE => () berghofe@43227: | SOME _ => error ("Cannot associate a type with " ^ s ^ berghofe@43227: "\nsince it is no record or enumeration type"); berghofe@43227: berghofe@43227: fun check_enum [] [] = NONE berghofe@43227: | check_enum els [] = SOME ("has no element(s) " ^ commas els) berghofe@43227: | check_enum [] cs = SOME ("has extra element(s) " ^ berghofe@43227: commas (map (Long_Name.base_name o fst) cs)) berghofe@43227: | check_enum (el :: els) ((cname, _) :: cs) = berghofe@43227: if lcase_eq (el, cname) then check_enum els cs berghofe@43227: else SOME ("either has no element " ^ el ^ berghofe@43227: " or it is at the wrong position"); berghofe@43227: berghofe@47596: fun invert_map [] = I berghofe@47596: | invert_map cmap = berghofe@47596: map (apfst (the o AList.lookup (op =) (map swap cmap))); berghofe@47596: berghofe@43267: fun add_type_def prfx (s, Basic_Type ty) (ids, thy) = berghofe@43267: (check_no_assoc thy prfx s; berghofe@43227: (ids, berghofe@43227: Typedecl.abbrev_global (Binding.name s, [], NoSyn) berghofe@43267: (mk_type thy prfx ty) thy |> snd)) berghofe@41809: berghofe@43267: | add_type_def prfx (s, Enum_Type els) ((tab, ctxt), thy) = berghofe@43227: let berghofe@43267: val (thy', tyname) = (case get_type thy prfx s of berghofe@43227: NONE => berghofe@43227: let berghofe@43227: val tyb = Binding.name s; berghofe@43227: val tyname = Sign.full_name thy tyb berghofe@43227: in berghofe@43227: (thy |> wenzelm@46572: Datatype.add_datatype {strict = true, quiet = true} berghofe@47596: [((tyb, [], NoSyn), berghofe@47596: map (fn s => (Binding.name s, [], NoSyn)) els)] |> snd |> berghofe@43227: add_enum_type s tyname, berghofe@43227: tyname) berghofe@43227: end berghofe@47596: | SOME (T as Type (tyname, []), cmap) => wenzelm@46777: (case Datatype.get_constrs thy tyname of berghofe@43227: NONE => assoc_ty_err thy T s "is not a datatype" berghofe@43267: | SOME cs => berghofe@47596: let val (prfx', _) = strip_prfx s berghofe@43267: in berghofe@47596: case check_enum (map (unprefix_pfun prfx') els) berghofe@47596: (invert_map cmap cs) of berghofe@43267: NONE => (thy, tyname) berghofe@43267: | SOME msg => assoc_ty_err thy T s msg berghofe@47438: end) berghofe@47596: | SOME (T, _) => assoc_ty_err thy T s "is not a datatype"); wenzelm@46777: val cs = map Const (the (Datatype.get_constrs thy' tyname)); berghofe@43227: in berghofe@43227: ((fold (Symtab.update_new o apsnd (rpair s)) (els ~~ cs) tab, berghofe@43227: fold Name.declare els ctxt), berghofe@43227: thy') berghofe@43227: end berghofe@41809: berghofe@43267: | add_type_def prfx (s, Array_Type (argtys, resty)) (ids, thy) = berghofe@43267: (check_no_assoc thy prfx s; berghofe@43227: (ids, berghofe@43227: Typedecl.abbrev_global (Binding.name s, [], NoSyn) berghofe@43267: (foldr1 HOLogic.mk_prodT (map (mk_type thy prfx) argtys) --> berghofe@43267: mk_type thy prfx resty) thy |> snd)) berghofe@41809: berghofe@43267: | add_type_def prfx (s, Record_Type fldtys) (ids, thy) = berghofe@41809: (ids, berghofe@43227: let val fldTs = maps (fn (flds, ty) => berghofe@43267: map (rpair (mk_type thy prfx ty)) flds) fldtys berghofe@43267: in case get_type thy prfx s of berghofe@43227: NONE => wenzelm@45524: Record.add_record ([], Binding.name s) NONE berghofe@43227: (map (fn (fld, T) => (Binding.name fld, T, NoSyn)) fldTs) thy berghofe@47596: | SOME (rT, cmap) => berghofe@43227: (case get_record_info thy rT of berghofe@43227: NONE => assoc_ty_err thy rT s "is not a record type" berghofe@43227: | SOME {fields, ...} => berghofe@47596: let val fields' = invert_map cmap fields berghofe@47596: in berghofe@47596: (case subtract (lcase_eq o pairself fst) fldTs fields' of berghofe@47596: [] => () berghofe@47596: | flds => assoc_ty_err thy rT s ("has extra field(s) " ^ berghofe@47596: commas (map (Long_Name.base_name o fst) flds)); berghofe@47596: map (fn (fld, T) => berghofe@47596: case AList.lookup lcase_eq fields' fld of berghofe@47596: NONE => assoc_ty_err thy rT s ("has no field " ^ fld) berghofe@47596: | SOME U => T = U orelse assoc_ty_err thy rT s berghofe@47596: ("has field " ^ berghofe@47596: fld ^ " whose type\n" ^ berghofe@47596: Syntax.string_of_typ_global thy U ^ berghofe@47596: "\ndoes not match declared type\n" ^ berghofe@47596: Syntax.string_of_typ_global thy T)) fldTs; berghofe@47596: thy) berghofe@47596: end) berghofe@43227: end) berghofe@41809: berghofe@43267: | add_type_def prfx (s, Pending_Type) (ids, thy) = berghofe@44901: (ids, berghofe@44901: case get_type thy prfx s of berghofe@44901: SOME _ => thy berghofe@44901: | NONE => Typedecl.typedecl_global berghofe@44901: (Binding.name s, [], NoSyn) thy |> snd); berghofe@41809: berghofe@41809: berghofe@43267: fun term_of_expr thy prfx types pfuns = berghofe@41809: let berghofe@41809: fun tm_of vs (Funct ("->", [e, e'])) = berghofe@41809: (HOLogic.mk_imp (fst (tm_of vs e), fst (tm_of vs e')), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("<->", [e, e'])) = berghofe@41809: (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("or", [e, e'])) = berghofe@41809: (HOLogic.mk_disj (fst (tm_of vs e), fst (tm_of vs e')), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("and", [e, e'])) = berghofe@41809: (HOLogic.mk_conj (fst (tm_of vs e), fst (tm_of vs e')), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("not", [e])) = berghofe@41809: (HOLogic.mk_not (fst (tm_of vs e)), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("=", [e, e'])) = berghofe@41809: (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("<>", [e, e'])) = (HOLogic.mk_not berghofe@41809: (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e'))), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("<", [e, e'])) = (HOLogic.mk_binrel @{const_name less} berghofe@41809: (fst (tm_of vs e), fst (tm_of vs e')), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct (">", [e, e'])) = (HOLogic.mk_binrel @{const_name less} berghofe@41809: (fst (tm_of vs e'), fst (tm_of vs e)), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("<=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq} berghofe@41809: (fst (tm_of vs e), fst (tm_of vs e')), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct (">=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq} berghofe@41809: (fst (tm_of vs e'), fst (tm_of vs e)), booleanN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("+", [e, e'])) = (HOLogic.mk_binop @{const_name plus} berghofe@41809: (fst (tm_of vs e), fst (tm_of vs e')), integerN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("-", [e, e'])) = (HOLogic.mk_binop @{const_name minus} berghofe@41809: (fst (tm_of vs e), fst (tm_of vs e')), integerN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("*", [e, e'])) = (HOLogic.mk_binop @{const_name times} berghofe@41809: (fst (tm_of vs e), fst (tm_of vs e')), integerN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("/", [e, e'])) = (HOLogic.mk_binop @{const_name divide} berghofe@41809: (fst (tm_of vs e), fst (tm_of vs e')), integerN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("div", [e, e'])) = (HOLogic.mk_binop @{const_name sdiv} berghofe@41809: (fst (tm_of vs e), fst (tm_of vs e')), integerN) berghofe@41809: berghofe@41883: | tm_of vs (Funct ("mod", [e, e'])) = (HOLogic.mk_binop @{const_name mod} berghofe@41809: (fst (tm_of vs e), fst (tm_of vs e')), integerN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("-", [e])) = berghofe@41809: (mk_unop @{const_name uminus} (fst (tm_of vs e)), integerN) berghofe@41809: berghofe@41809: | tm_of vs (Funct ("**", [e, e'])) = berghofe@41809: (Const (@{const_name power}, HOLogic.intT --> HOLogic.natT --> berghofe@41809: HOLogic.intT) $ fst (tm_of vs e) $ berghofe@41809: (@{const nat} $ fst (tm_of vs e')), integerN) berghofe@41809: berghofe@41809: | tm_of (tab, _) (Ident s) = berghofe@41809: (case Symtab.lookup tab s of berghofe@41809: SOME t_ty => t_ty berghofe@43370: | NONE => (case lookup_prfx prfx pfuns s of berghofe@43370: SOME (SOME ([], resty), t) => (t, resty) berghofe@43370: | _ => error ("Undeclared identifier " ^ s))) berghofe@41809: berghofe@41809: | tm_of _ (Number i) = (HOLogic.mk_number HOLogic.intT i, integerN) berghofe@41809: berghofe@41809: | tm_of vs (Quantifier (s, xs, ty, e)) = berghofe@41809: let berghofe@43267: val (ys, vs') = mk_variables thy prfx xs ty vs; berghofe@41809: val q = (case s of berghofe@41809: "for_all" => HOLogic.mk_all berghofe@41809: | "for_some" => HOLogic.mk_exists) berghofe@41809: in berghofe@41809: (fold_rev (fn Free (x, T) => fn t => q (x, T, t)) berghofe@41809: ys (fst (tm_of vs' e)), berghofe@41809: booleanN) berghofe@41809: end berghofe@41809: berghofe@41809: | tm_of vs (Funct (s, es)) = berghofe@41809: berghofe@41809: (* record field selection *) berghofe@41809: (case try (unprefix "fld_") s of berghofe@41809: SOME fname => (case es of berghofe@41809: [e] => berghofe@43227: let berghofe@43227: val (t, rcdty) = tm_of vs e; berghofe@47596: val (rT, cmap) = mk_type' thy prfx rcdty berghofe@43227: in case (get_record_info thy rT, lookup types rcdty) of berghofe@43227: (SOME {fields, ...}, SOME (Record_Type fldtys)) => berghofe@47596: (case (find_field cmap fname fields, berghofe@43227: find_field' fname fldtys) of berghofe@43227: (SOME (fname', fT), SOME fldty) => berghofe@43227: (Const (fname', rT --> fT) $ t, fldty) berghofe@43227: | _ => error ("Record " ^ rcdty ^ berghofe@41809: " has no field named " ^ fname)) berghofe@41809: | _ => error (rcdty ^ " is not a record type") berghofe@41809: end berghofe@41809: | _ => error ("Function " ^ s ^ " expects one argument")) berghofe@41809: | NONE => berghofe@41809: berghofe@41809: (* record field update *) berghofe@41809: (case try (unprefix "upf_") s of berghofe@41809: SOME fname => (case es of berghofe@41809: [e, e'] => berghofe@41809: let berghofe@41809: val (t, rcdty) = tm_of vs e; berghofe@47596: val (rT, cmap) = mk_type' thy prfx rcdty; berghofe@41809: val (u, fldty) = tm_of vs e'; berghofe@43267: val fT = mk_type thy prfx fldty berghofe@43227: in case get_record_info thy rT of berghofe@43227: SOME {fields, ...} => berghofe@47596: (case find_field cmap fname fields of berghofe@43227: SOME (fname', fU) => berghofe@43227: if fT = fU then berghofe@43227: (Const (fname' ^ "_update", berghofe@41809: (fT --> fT) --> rT --> rT) $ berghofe@41809: Abs ("x", fT, u) $ t, berghofe@41809: rcdty) berghofe@43227: else error ("Type\n" ^ berghofe@43227: Syntax.string_of_typ_global thy fT ^ berghofe@43227: "\ndoes not match type\n" ^ berghofe@43227: Syntax.string_of_typ_global thy fU ^ berghofe@43227: "\nof field " ^ fname) berghofe@41809: | NONE => error ("Record " ^ rcdty ^ berghofe@41809: " has no field named " ^ fname)) berghofe@41809: | _ => error (rcdty ^ " is not a record type") berghofe@41809: end berghofe@41809: | _ => error ("Function " ^ s ^ " expects two arguments")) berghofe@41809: | NONE => berghofe@41809: berghofe@41809: (* enumeration type to integer *) berghofe@41809: (case try (unsuffix "__pos") s of berghofe@41809: SOME tyname => (case es of berghofe@41809: [e] => (Const (@{const_name pos}, berghofe@43267: mk_type thy prfx tyname --> HOLogic.intT) $ fst (tm_of vs e), berghofe@43267: integerN) berghofe@41809: | _ => error ("Function " ^ s ^ " expects one argument")) berghofe@41809: | NONE => berghofe@41809: berghofe@41809: (* integer to enumeration type *) berghofe@41809: (case try (unsuffix "__val") s of berghofe@41809: SOME tyname => (case es of berghofe@41809: [e] => (Const (@{const_name val}, berghofe@43267: HOLogic.intT --> mk_type thy prfx tyname) $ fst (tm_of vs e), berghofe@43267: tyname) berghofe@41809: | _ => error ("Function " ^ s ^ " expects one argument")) berghofe@41809: | NONE => berghofe@41809: berghofe@41809: (* successor / predecessor of enumeration type element *) berghofe@41809: if s = "succ" orelse s = "pred" then (case es of berghofe@41809: [e] => berghofe@41809: let berghofe@41809: val (t, tyname) = tm_of vs e; berghofe@43267: val T = mk_type thy prfx tyname berghofe@41809: in (Const berghofe@41809: (if s = "succ" then @{const_name succ} berghofe@41809: else @{const_name pred}, T --> T) $ t, tyname) berghofe@41809: end berghofe@41809: | _ => error ("Function " ^ s ^ " expects one argument")) berghofe@41809: berghofe@41809: (* user-defined proof function *) berghofe@41809: else berghofe@43267: (case lookup_prfx prfx pfuns s of berghofe@41809: SOME (SOME (_, resty), t) => berghofe@41809: (list_comb (t, map (fst o tm_of vs) es), resty) berghofe@41809: | _ => error ("Undeclared proof function " ^ s)))))) berghofe@41809: berghofe@41809: | tm_of vs (Element (e, es)) = berghofe@41809: let val (t, ty) = tm_of vs e berghofe@41809: in case lookup types ty of berghofe@41809: SOME (Array_Type (_, elty)) => berghofe@41809: (t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es), elty) berghofe@41809: | _ => error (ty ^ " is not an array type") berghofe@41809: end berghofe@41809: berghofe@41809: | tm_of vs (Update (e, es, e')) = berghofe@41809: let val (t, ty) = tm_of vs e berghofe@41809: in case lookup types ty of berghofe@41809: SOME (Array_Type (idxtys, elty)) => berghofe@41809: let berghofe@43267: val T = foldr1 HOLogic.mk_prodT berghofe@43267: (map (mk_type thy prfx) idxtys); berghofe@43267: val U = mk_type thy prfx elty; berghofe@41809: val fT = T --> U berghofe@41809: in berghofe@41809: (Const (@{const_name fun_upd}, fT --> T --> U --> fT) $ berghofe@41809: t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es) $ berghofe@41809: fst (tm_of vs e'), berghofe@41809: ty) berghofe@41809: end berghofe@41809: | _ => error (ty ^ " is not an array type") berghofe@41809: end berghofe@41809: berghofe@41809: | tm_of vs (Record (s, flds)) = berghofe@43227: let berghofe@47596: val (T, cmap) = mk_type' thy prfx s; berghofe@43227: val {extension = (ext_name, _), fields, ...} = berghofe@43227: (case get_record_info thy T of berghofe@43227: NONE => error (s ^ " is not a record type") berghofe@43227: | SOME info => info); berghofe@43227: val flds' = map (apsnd (tm_of vs)) flds; berghofe@47596: val fnames = fields |> invert_map cmap |> berghofe@47596: map (Long_Name.base_name o fst); berghofe@43227: val fnames' = map fst flds; berghofe@43227: val (fvals, ftys) = split_list (map (fn s' => berghofe@43227: case AList.lookup lcase_eq flds' s' of berghofe@43227: SOME fval_ty => fval_ty berghofe@43227: | NONE => error ("Field " ^ s' ^ " missing in record " ^ s)) berghofe@43227: fnames); berghofe@43227: val _ = (case subtract lcase_eq fnames fnames' of berghofe@43227: [] => () berghofe@43227: | xs => error ("Extra field(s) " ^ commas xs ^ berghofe@43227: " in record " ^ s)); berghofe@43227: val _ = (case duplicates (op =) fnames' of berghofe@43227: [] => () berghofe@43227: | xs => error ("Duplicate field(s) " ^ commas xs ^ berghofe@43227: " in record " ^ s)) berghofe@43227: in berghofe@43227: (list_comb berghofe@43227: (Const (ext_name, berghofe@43267: map (mk_type thy prfx) ftys @ [HOLogic.unitT] ---> T), berghofe@43227: fvals @ [HOLogic.unit]), berghofe@43227: s) berghofe@43227: end berghofe@41809: berghofe@41809: | tm_of vs (Array (s, default, assocs)) = berghofe@41809: (case lookup types s of berghofe@41809: SOME (Array_Type (idxtys, elty)) => berghofe@41809: let berghofe@43267: val Ts = map (mk_type thy prfx) idxtys; berghofe@41809: val T = foldr1 HOLogic.mk_prodT Ts; berghofe@43267: val U = mk_type thy prfx elty; berghofe@41809: fun mk_idx' T (e, NONE) = HOLogic.mk_set T [fst (tm_of vs e)] berghofe@41809: | mk_idx' T (e, SOME e') = Const (@{const_name atLeastAtMost}, berghofe@41809: T --> T --> HOLogic.mk_setT T) $ berghofe@41809: fst (tm_of vs e) $ fst (tm_of vs e'); berghofe@41809: fun mk_idx idx = berghofe@41809: if length Ts <> length idx then berghofe@41809: error ("Arity mismatch in construction of array " ^ s) berghofe@41809: else foldr1 mk_times (map2 mk_idx' Ts idx); berghofe@41809: fun mk_upd (idxs, e) t = berghofe@41809: if length idxs = 1 andalso forall (is_none o snd) (hd idxs) berghofe@41809: then berghofe@41809: Const (@{const_name fun_upd}, (T --> U) --> berghofe@41809: T --> U --> T --> U) $ t $ berghofe@41809: foldl1 HOLogic.mk_prod berghofe@41809: (map (fst o tm_of vs o fst) (hd idxs)) $ berghofe@41809: fst (tm_of vs e) berghofe@41809: else berghofe@41809: Const (@{const_name fun_upds}, (T --> U) --> berghofe@41809: HOLogic.mk_setT T --> U --> T --> U) $ t $ berghofe@41809: foldl1 (HOLogic.mk_binop @{const_name sup}) berghofe@41809: (map mk_idx idxs) $ berghofe@41809: fst (tm_of vs e) berghofe@41809: in berghofe@41809: (fold mk_upd assocs berghofe@41809: (case default of berghofe@41809: SOME e => Abs ("x", T, fst (tm_of vs e)) berghofe@41809: | NONE => Const (@{const_name undefined}, T --> U)), berghofe@41809: s) berghofe@41809: end berghofe@41809: | _ => error (s ^ " is not an array type")) berghofe@41809: berghofe@41809: in tm_of end; berghofe@41809: berghofe@41809: berghofe@43267: fun term_of_rule thy prfx types pfuns ids rule = berghofe@43267: let val tm_of = fst o term_of_expr thy prfx types pfuns ids berghofe@41809: in case rule of berghofe@41809: Inference_Rule (es, e) => Logic.list_implies berghofe@41809: (map (HOLogic.mk_Trueprop o tm_of) es, HOLogic.mk_Trueprop (tm_of e)) berghofe@41809: | Substitution_Rule (es, e, e') => Logic.list_implies berghofe@41809: (map (HOLogic.mk_Trueprop o tm_of) es, berghofe@41809: HOLogic.mk_Trueprop (HOLogic.mk_eq (tm_of e, tm_of e'))) berghofe@41809: end; berghofe@41809: berghofe@41809: berghofe@41809: val builtin = Symtab.make (map (rpair ()) berghofe@41809: ["->", "<->", "or", "and", "not", "=", "<>", "<", ">", "<=", ">=", berghofe@41809: "+", "-", "*", "/", "div", "mod", "**"]); berghofe@41809: berghofe@41809: fun complex_expr (Number _) = false berghofe@41809: | complex_expr (Ident _) = false berghofe@41809: | complex_expr (Funct (s, es)) = berghofe@41809: not (Symtab.defined builtin s) orelse exists complex_expr es berghofe@41809: | complex_expr (Quantifier (_, _, _, e)) = complex_expr e berghofe@41809: | complex_expr _ = true; berghofe@41809: berghofe@41809: fun complex_rule (Inference_Rule (es, e)) = berghofe@41809: complex_expr e orelse exists complex_expr es berghofe@41809: | complex_rule (Substitution_Rule (es, e, e')) = berghofe@41809: complex_expr e orelse complex_expr e' orelse berghofe@41809: exists complex_expr es; berghofe@41809: berghofe@41809: val is_pfun = berghofe@41809: Symtab.defined builtin orf berghofe@41809: can (unprefix "fld_") orf can (unprefix "upf_") orf berghofe@41809: can (unsuffix "__pos") orf can (unsuffix "__val") orf berghofe@41809: equal "succ" orf equal "pred"; berghofe@41809: berghofe@41809: fun fold_opt f = the_default I o Option.map f; berghofe@41809: fun fold_pair f g (x, y) = f x #> g y; berghofe@41809: berghofe@41809: fun fold_expr f g (Funct (s, es)) = f s #> fold (fold_expr f g) es berghofe@41809: | fold_expr f g (Ident s) = g s berghofe@41809: | fold_expr f g (Number _) = I berghofe@41809: | fold_expr f g (Quantifier (_, _, _, e)) = fold_expr f g e berghofe@41809: | fold_expr f g (Element (e, es)) = berghofe@41809: fold_expr f g e #> fold (fold_expr f g) es berghofe@41809: | fold_expr f g (Update (e, es, e')) = berghofe@41809: fold_expr f g e #> fold (fold_expr f g) es #> fold_expr f g e' berghofe@41809: | fold_expr f g (Record (_, flds)) = fold (fold_expr f g o snd) flds berghofe@41809: | fold_expr f g (Array (_, default, assocs)) = berghofe@41809: fold_opt (fold_expr f g) default #> berghofe@41809: fold (fold_pair berghofe@41809: (fold (fold (fold_pair berghofe@41809: (fold_expr f g) (fold_opt (fold_expr f g))))) berghofe@41809: (fold_expr f g)) assocs; berghofe@41809: berghofe@43370: fun add_expr_pfuns funs = fold_expr berghofe@43370: (fn s => if is_pfun s then I else insert (op =) s) berghofe@43370: (fn s => if is_none (lookup funs s) then I else insert (op =) s); berghofe@41809: berghofe@41809: val add_expr_idents = fold_expr (K I) (insert (op =)); berghofe@41809: berghofe@43267: fun pfun_type thy prfx (argtys, resty) = berghofe@43267: map (mk_type thy prfx) argtys ---> mk_type thy prfx resty; berghofe@41809: berghofe@43267: fun check_pfun_type thy prfx s t optty1 optty2 = berghofe@41809: let berghofe@41809: val T = fastype_of t; berghofe@41809: fun check ty = berghofe@43267: let val U = pfun_type thy prfx ty berghofe@41809: in berghofe@41809: T = U orelse berghofe@41809: error ("Type\n" ^ berghofe@41809: Syntax.string_of_typ_global thy T ^ berghofe@41809: "\nof function " ^ berghofe@41809: Syntax.string_of_term_global thy t ^ berghofe@41809: " associated with proof function " ^ s ^ berghofe@41809: "\ndoes not match declared type\n" ^ berghofe@41809: Syntax.string_of_typ_global thy U) berghofe@41809: end berghofe@41809: in (Option.map check optty1; Option.map check optty2; ()) end; berghofe@41809: berghofe@41809: fun upd_option x y = if is_some x then x else y; berghofe@41809: berghofe@43267: fun check_pfuns_types thy prfx funs = berghofe@41809: Symtab.map (fn s => fn (optty, t) => berghofe@47438: let val optty' = lookup funs (unprefix_pfun prfx s) berghofe@41809: in berghofe@43267: (check_pfun_type thy prfx s t optty optty'; berghofe@41809: (NONE |> upd_option optty |> upd_option optty', t)) berghofe@41809: end); berghofe@41809: berghofe@41809: berghofe@41809: (** the VC store **) berghofe@41809: berghofe@49182: fun pp_vcs msg vcs = Pretty.big_list msg (map (Pretty.str o fst) vcs); berghofe@49182: berghofe@49182: fun pp_open_vcs [] = Pretty.str "All verification conditions have been proved." berghofe@49182: | pp_open_vcs vcs = pp_vcs berghofe@49182: "The following verification conditions remain to be proved:" vcs; berghofe@49182: berghofe@49182: fun partition_vcs vcs = VCtab.fold_rev berghofe@49182: (fn (name, (trace, SOME thms, ps, cs)) => berghofe@49182: apfst (cons (name, (trace, thms, ps, cs))) berghofe@49182: | (name, (trace, NONE, ps, cs)) => berghofe@49182: apsnd (cons (name, (trace, ps, cs)))) berghofe@49182: vcs ([], []); berghofe@49182: berghofe@49182: fun insert_break prt = Pretty.blk (0, [Pretty.fbrk, prt]); berghofe@49182: berghofe@49182: fun print_open_vcs f vcs = berghofe@49182: (Pretty.writeln (f (pp_open_vcs (snd (partition_vcs vcs)))); vcs); berghofe@41809: berghofe@47438: fun set_env ctxt defs types funs ids vcs path prefix thy = VCs.map (fn berghofe@43227: {pfuns, type_map, env = NONE} => berghofe@47438: {pfuns = pfuns, berghofe@43227: type_map = type_map, berghofe@47438: env = SOME berghofe@47438: {ctxt = ctxt, defs = defs, types = types, funs = funs, berghofe@47438: pfuns = check_pfuns_types thy prefix funs pfuns, berghofe@49182: ids = ids, proving = false, vcs = print_open_vcs I vcs, path = path, berghofe@47438: prefix = prefix}} berghofe@41809: | _ => err_unfinished ()) thy; berghofe@41809: berghofe@41809: fun mk_pat s = (case Int.fromString s of berghofe@41809: SOME i => [HOLogic.mk_Trueprop (Var (("C", i), HOLogic.boolT))] berghofe@41809: | NONE => error ("Bad conclusion identifier: C" ^ s)); berghofe@41809: berghofe@49182: fun mk_vc thy prfx types pfuns ids name_concl (tr, proved, ps, cs) = berghofe@41809: let val prop_of = berghofe@43267: HOLogic.mk_Trueprop o fst o term_of_expr thy prfx types pfuns ids berghofe@41809: in berghofe@41809: (tr, proved, berghofe@41809: Element.Assumes (map (fn (s', e) => berghofe@41809: ((Binding.name ("H" ^ s'), []), [(prop_of e, [])])) ps), berghofe@41809: Element.Shows (map (fn (s', e) => berghofe@49182: (if name_concl then (Binding.name ("C" ^ s'), []) berghofe@49182: else Attrib.empty_binding, berghofe@49182: [(prop_of e, mk_pat s')])) cs)) berghofe@41809: end; berghofe@41809: berghofe@41809: fun fold_vcs f vcs = berghofe@41809: VCtab.fold (fn (_, (_, _, ps, cs)) => fold f ps #> fold f cs) vcs; berghofe@41809: berghofe@43370: fun pfuns_of_vcs prfx funs pfuns vcs = berghofe@43370: fold_vcs (add_expr_pfuns funs o snd) vcs [] |> berghofe@43267: filter (is_none o lookup_prfx prfx pfuns); berghofe@41809: berghofe@43267: fun declare_missing_pfuns thy prfx funs pfuns vcs (tab, ctxt) = berghofe@41809: let berghofe@41809: val (fs, (tys, Ts)) = berghofe@43370: pfuns_of_vcs prfx funs pfuns vcs |> berghofe@41809: map_filter (fn s => lookup funs s |> berghofe@43267: Option.map (fn ty => (s, (SOME ty, pfun_type thy prfx ty)))) |> berghofe@41809: split_list ||> split_list; wenzelm@44208: val (fs', ctxt') = fold_map Name.variant fs ctxt berghofe@41809: in berghofe@41809: (fold Symtab.update_new (fs ~~ (tys ~~ map Free (fs' ~~ Ts))) pfuns, berghofe@41809: Element.Fixes (map2 (fn s => fn T => berghofe@41809: (Binding.name s, SOME T, NoSyn)) fs' Ts), berghofe@41809: (tab, ctxt')) berghofe@41809: end; berghofe@41809: berghofe@47438: fun map_pfuns f {pfuns, type_map, env} = berghofe@47438: {pfuns = f pfuns, type_map = type_map, env = env} berghofe@47438: berghofe@47438: fun map_pfuns_env f {pfuns, type_map, berghofe@47438: env = SOME {ctxt, defs, types, funs, pfuns = pfuns_env, berghofe@47438: ids, proving, vcs, path, prefix}} = berghofe@47438: if proving then err_unfinished () berghofe@47438: else berghofe@47438: {pfuns = pfuns, type_map = type_map, berghofe@47438: env = SOME {ctxt = ctxt, defs = defs, types = types, funs = funs, berghofe@47438: pfuns = f pfuns_env, ids = ids, proving = false, vcs = vcs, berghofe@47438: path = path, prefix = prefix}}; berghofe@47438: berghofe@41809: fun add_proof_fun prep (s, (optty, raw_t)) thy = berghofe@47438: VCs.map (fn data as {env, ...} => berghofe@47438: let berghofe@47438: val (optty', prfx, map_pf) = (case env of berghofe@47438: SOME {funs, prefix, ...} => berghofe@47438: (lookup funs (unprefix_pfun prefix s), berghofe@47438: prefix, map_pfuns_env) berghofe@47438: | NONE => (NONE, "", map_pfuns)); berghofe@47438: val optty'' = NONE |> upd_option optty |> upd_option optty'; berghofe@47438: val t = prep (Option.map (pfun_type thy prfx) optty'') raw_t; berghofe@47438: val _ = (case fold_aterms (fn u => berghofe@47438: if is_Var u orelse is_Free u then insert (op =) u else I) t [] of berghofe@47438: [] => () berghofe@47438: | ts => error ("Term\n" ^ Syntax.string_of_term_global thy t ^ berghofe@47438: "\nto be associated with proof function " ^ s ^ berghofe@47438: " contains free variable(s) " ^ berghofe@47438: commas (map (Syntax.string_of_term_global thy) ts))); berghofe@47438: in berghofe@47438: (check_pfun_type thy prfx s t optty optty'; berghofe@47438: if is_some optty'' orelse is_none env then berghofe@47438: map_pf (Symtab.update_new (s, (optty'', t))) data berghofe@47438: handle Symtab.DUP _ => error ("Proof function " ^ s ^ berghofe@47438: " already associated with function") berghofe@47438: else error ("Undeclared proof function " ^ s)) berghofe@47438: end) thy; berghofe@41809: berghofe@47596: fun check_mapping _ _ [] _ = () berghofe@47596: | check_mapping err s cmap cs = berghofe@47596: (case duplicates (op = o pairself fst) cmap of berghofe@47596: [] => (case duplicates (op = o pairself snd) cmap of berghofe@47596: [] => (case subtract (op = o apsnd snd) cs cmap of berghofe@47596: [] => (case subtract (op = o apfst snd) cmap cs of berghofe@47596: [] => () berghofe@47596: | zs => err ("has extra " ^ s ^ "(s) " ^ commas zs)) berghofe@47596: | zs => err ("does not have " ^ s ^ "(s) " ^ berghofe@47596: commas (map snd zs))) berghofe@47596: | zs => error ("Several SPARK names are mapped to " ^ berghofe@47596: commas (map snd zs))) berghofe@47596: | zs => error ("The SPARK names " ^ commas (map fst zs) ^ berghofe@47596: " are mapped to more than one name")); berghofe@47596: berghofe@47596: fun add_type (s, (T as Type (tyname, Ts), cmap)) thy = berghofe@47596: let val cmap' = map (apsnd (Sign.intern_const thy)) cmap berghofe@47596: in berghofe@47596: thy |> berghofe@47596: VCs.map (fn berghofe@47596: {env = SOME _, ...} => err_unfinished () berghofe@47596: | {pfuns, type_map, env} => berghofe@47596: {pfuns = pfuns, berghofe@47596: type_map = Symtab.update_new (s, (T, cmap')) type_map, berghofe@47596: env = env} berghofe@47596: handle Symtab.DUP _ => error ("SPARK type " ^ s ^ berghofe@47596: " already associated with type")) |> berghofe@47596: (fn thy' => berghofe@47596: case Datatype.get_constrs thy' tyname of berghofe@47596: NONE => (case get_record_info thy' T of berghofe@47596: NONE => thy' berghofe@47596: | SOME {fields, ...} => berghofe@47596: (check_mapping (assoc_ty_err thy' T s) "field" berghofe@47596: cmap' (map fst fields); berghofe@47596: thy')) berghofe@47596: | SOME cs => berghofe@47596: if null Ts then berghofe@47596: (map berghofe@47596: (fn (_, Type (_, [])) => () berghofe@47596: | (cname, _) => assoc_ty_err thy' T s berghofe@47596: ("has illegal constructor " ^ berghofe@47596: Long_Name.base_name cname)) cs; berghofe@47596: check_mapping (assoc_ty_err thy' T s) "constructor" berghofe@47596: cmap' (map fst cs); berghofe@47596: add_enum_type s tyname thy') berghofe@47596: else assoc_ty_err thy' T s "is illegal") berghofe@47596: end berghofe@47596: | add_type (s, (T, _)) thy = assoc_ty_err thy T s "is illegal"; berghofe@43227: berghofe@41809: val is_closed = is_none o #env o VCs.get; berghofe@41809: berghofe@49182: fun lookup_vc thy name_concl name = berghofe@41809: (case VCs.get thy of berghofe@47438: {env = SOME {vcs, types, funs, pfuns, ids, ctxt, prefix, ...}, ...} => berghofe@41809: (case VCtab.lookup vcs name of berghofe@43267: SOME vc => berghofe@41809: let val (pfuns', ctxt', ids') = berghofe@43267: declare_missing_pfuns thy prefix funs pfuns vcs ids berghofe@49182: in berghofe@49182: SOME (ctxt @ [ctxt'], berghofe@49182: mk_vc thy prefix types pfuns' ids' name_concl vc) berghofe@49182: end berghofe@41809: | NONE => NONE) berghofe@41809: | _ => NONE); berghofe@41809: berghofe@49182: fun get_vcs thy name_concl = (case VCs.get thy of berghofe@47438: {env = SOME {vcs, types, funs, pfuns, ids, ctxt, defs, prefix, ...}, ...} => berghofe@41809: let val (pfuns', ctxt', ids') = berghofe@43267: declare_missing_pfuns thy prefix funs pfuns vcs ids berghofe@41809: in berghofe@41809: (ctxt @ [ctxt'], defs, berghofe@41809: VCtab.dest vcs |> berghofe@49182: map (apsnd (mk_vc thy prefix types pfuns' ids' name_concl))) berghofe@41809: end berghofe@41809: | _ => ([], [], [])); berghofe@41809: berghofe@42767: fun mark_proved name thms = VCs.map (fn berghofe@43227: {pfuns, type_map, berghofe@47438: env = SOME {ctxt, defs, types, funs, pfuns = pfuns_env, berghofe@47438: ids, vcs, path, prefix, ...}} => berghofe@41809: {pfuns = pfuns, berghofe@43227: type_map = type_map, berghofe@41809: env = SOME {ctxt = ctxt, defs = defs, berghofe@47438: types = types, funs = funs, pfuns = pfuns_env, berghofe@47438: ids = ids, berghofe@41809: proving = true, berghofe@49182: vcs = print_open_vcs insert_break (VCtab.map_entry name berghofe@49182: (fn (trace, _, ps, cs) => (trace, SOME thms, ps, cs)) vcs), berghofe@43267: path = path, berghofe@43267: prefix = prefix}} berghofe@41809: | x => x); berghofe@41809: berghofe@49182: fun close incomplete thy = berghofe@43227: thy |> berghofe@43227: VCs.map (fn berghofe@43227: {pfuns, type_map, env = SOME {vcs, path, ...}} => berghofe@49182: let berghofe@49182: val (proved, unproved) = partition_vcs vcs; berghofe@49182: val _ = Thm.join_proofs (maps (#2 o snd) proved); berghofe@49182: val (proved', proved'') = List.partition (fn (_, (_, thms, _, _)) => berghofe@49182: exists (#oracle o Thm.status_of) thms) proved berghofe@49182: in berghofe@49182: (if null unproved then () berghofe@49182: else (if incomplete then warning else error) berghofe@49182: (Pretty.string_of (pp_open_vcs unproved)); berghofe@49182: if null proved' then () berghofe@49182: else warning (Pretty.string_of (pp_vcs berghofe@49182: "The following VCs are not marked as proved \ berghofe@49182: \because their proofs contain oracles:" proved')); berghofe@49182: File.write (Path.ext "prv" path) berghofe@49182: (implode (map (fn (s, _) => snd (strip_number s) ^ berghofe@49182: " -- proved by " ^ Distribution.version ^ "\n") proved'')); berghofe@49182: {pfuns = pfuns, type_map = type_map, env = NONE}) berghofe@49182: end berghofe@43227: | _ => error "No SPARK environment is currently open") |> berghofe@43227: Sign.parent_path; berghofe@41809: berghofe@41809: berghofe@41809: (** set up verification conditions **) berghofe@41809: berghofe@41809: fun partition_opt f = berghofe@41809: let berghofe@41809: fun part ys zs [] = (rev ys, rev zs) berghofe@41809: | part ys zs (x :: xs) = (case f x of berghofe@41809: SOME y => part (y :: ys) zs xs berghofe@41809: | NONE => part ys (x :: zs) xs) berghofe@41809: in part [] [] end; berghofe@41809: berghofe@41809: fun dest_def (id, (Substitution_Rule ([], Ident s, rhs))) = SOME (id, (s, rhs)) berghofe@41809: | dest_def _ = NONE; berghofe@41809: berghofe@41809: fun mk_rulename (s, i) = Binding.name (s ^ string_of_int i); berghofe@41809: berghofe@43267: fun add_const prfx (s, ty) ((tab, ctxt), thy) = berghofe@41809: let berghofe@43267: val T = mk_type thy prfx ty; berghofe@41809: val b = Binding.name s; berghofe@41809: val c = Const (Sign.full_name thy b, T) berghofe@41809: in berghofe@41809: (c, berghofe@41809: ((Symtab.update (s, (c, ty)) tab, Name.declare s ctxt), berghofe@41809: Sign.add_consts_i [(b, T, NoSyn)] thy)) berghofe@41809: end; berghofe@41809: berghofe@43267: fun add_def prfx types pfuns consts (id, (s, e)) (ids as (tab, ctxt), thy) = berghofe@41809: (case lookup consts s of berghofe@41809: SOME ty => berghofe@41809: let berghofe@43267: val (t, ty') = term_of_expr thy prfx types pfuns ids e; berghofe@43267: val T = mk_type thy prfx ty; berghofe@43267: val T' = mk_type thy prfx ty'; berghofe@42749: val _ = T = T' orelse berghofe@41809: error ("Declared type " ^ ty ^ " of " ^ s ^ berghofe@41809: "\ndoes not match type " ^ ty' ^ " in definition"); berghofe@41809: val id' = mk_rulename id; berghofe@41809: val lthy = Named_Target.theory_init thy; berghofe@41809: val ((t', (_, th)), lthy') = Specification.definition berghofe@41809: (NONE, ((id', []), HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@42749: (Free (s, T), t)))) lthy; wenzelm@43232: val phi = Proof_Context.export_morphism lthy' lthy berghofe@41809: in berghofe@41809: ((id', Morphism.thm phi th), berghofe@41809: ((Symtab.update (s, (Morphism.term phi t', ty)) tab, berghofe@41809: Name.declare s ctxt), berghofe@41809: Local_Theory.exit_global lthy')) berghofe@41809: end berghofe@41809: | NONE => error ("Undeclared constant " ^ s)); berghofe@41809: berghofe@43267: fun add_var prfx (s, ty) (ids, thy) = berghofe@43267: let val ([Free p], ids') = mk_variables thy prfx [s] ty ids berghofe@41809: in (p, (ids', thy)) end; berghofe@41809: berghofe@43267: fun add_init_vars prfx vcs (ids_thy as ((tab, _), _)) = berghofe@43267: fold_map (add_var prfx) berghofe@41809: (map_filter berghofe@41809: (fn s => case try (unsuffix "~") s of berghofe@41809: SOME s' => (case Symtab.lookup tab s' of berghofe@41809: SOME (_, ty) => SOME (s, ty) berghofe@41809: | NONE => error ("Undeclared identifier " ^ s')) berghofe@41809: | NONE => NONE) berghofe@41809: (fold_vcs (add_expr_idents o snd) vcs [])) berghofe@41809: ids_thy; berghofe@41809: berghofe@41809: fun is_trivial_vc ([], [(_, Ident "true")]) = true berghofe@41809: | is_trivial_vc _ = false; berghofe@41809: berghofe@41809: fun rulenames rules = commas berghofe@41809: (map (fn ((s, i), _) => s ^ "(" ^ string_of_int i ^ ")") rules); berghofe@41809: berghofe@41809: (* sort definitions according to their dependency *) berghofe@43370: fun sort_defs _ _ _ _ [] sdefs = rev sdefs berghofe@43370: | sort_defs prfx funs pfuns consts defs sdefs = berghofe@41809: (case find_first (fn (_, (_, e)) => berghofe@43370: forall (is_some o lookup_prfx prfx pfuns) berghofe@43370: (add_expr_pfuns funs e []) andalso berghofe@41809: forall (fn id => berghofe@41809: member (fn (s, (_, (s', _))) => s = s') sdefs id orelse berghofe@42749: consts id) berghofe@41809: (add_expr_idents e [])) defs of berghofe@43370: SOME d => sort_defs prfx funs pfuns consts berghofe@41809: (remove (op =) d defs) (d :: sdefs) berghofe@41809: | NONE => error ("Bad definitions: " ^ rulenames defs)); berghofe@41809: berghofe@43227: fun set_vcs ({types, vars, consts, funs} : decls) berghofe@49468: (rules, _) ((_, ident), vcs) path opt_prfx thy = berghofe@41809: let berghofe@49468: val prfx' = berghofe@49468: if opt_prfx = "" then berghofe@49468: space_implode "__" (Long_Name.explode (Long_Name.qualifier ident)) berghofe@49468: else opt_prfx; berghofe@49468: val prfx = to_lower prfx'; berghofe@41809: val {pfuns, ...} = VCs.get thy; berghofe@42749: val (defs, rules') = partition_opt dest_def rules; berghofe@41809: val consts' = berghofe@42749: subtract (fn ((_, (s, _)), (s', _)) => s = s') defs (items consts); berghofe@41809: (* ignore all complex rules in rls files *) berghofe@41809: val (rules'', other_rules) = berghofe@41809: List.partition (complex_rule o snd) rules'; berghofe@41809: val _ = if null rules'' then () berghofe@41809: else warning ("Ignoring rules: " ^ rulenames rules''); berghofe@41809: berghofe@41809: val vcs' = VCtab.make (maps (fn (tr, vcs) => berghofe@42767: map (fn (s, (ps, cs)) => (s, (tr, NONE, ps, cs))) berghofe@41809: (filter_out (is_trivial_vc o snd) vcs)) vcs); berghofe@41809: berghofe@41809: val _ = (case filter_out (is_some o lookup funs) berghofe@43370: (pfuns_of_vcs prfx funs pfuns vcs') of berghofe@41809: [] => () berghofe@41809: | fs => error ("Undeclared proof function(s) " ^ commas fs)); berghofe@41809: berghofe@41809: val (((defs', vars''), ivars), (ids, thy')) = berghofe@41809: ((Symtab.empty |> wenzelm@46611: Symtab.update ("false", (@{term False}, booleanN)) |> wenzelm@46611: Symtab.update ("true", (@{term True}, booleanN)), berghofe@43227: Name.context), berghofe@47595: thy |> Sign.add_path berghofe@49468: ((if prfx' = "" then "" else prfx' ^ "__") ^ Long_Name.base_name ident)) |> berghofe@43267: fold (add_type_def prfx) (items types) |> berghofe@43267: fold (snd oo add_const prfx) consts' |> (fn ids_thy as ((tab, _), _) => berghofe@42749: ids_thy |> berghofe@43267: fold_map (add_def prfx types pfuns consts) berghofe@43370: (sort_defs prfx funs pfuns (Symtab.defined tab) defs []) ||>> berghofe@43267: fold_map (add_var prfx) (items vars) ||>> berghofe@43267: add_init_vars prfx vcs'); berghofe@41809: berghofe@41809: val ctxt = berghofe@41809: [Element.Fixes (map (fn (s, T) => berghofe@41809: (Binding.name s, SOME T, NoSyn)) (vars'' @ ivars)), berghofe@41809: Element.Assumes (map (fn (id, rl) => berghofe@41809: ((mk_rulename id, []), berghofe@43267: [(term_of_rule thy' prfx types pfuns ids rl, [])])) berghofe@41809: other_rules), wenzelm@43311: Element.Notes ("", [((Binding.name "defns", []), map (rpair [] o single o snd) defs')])] berghofe@41809: berghofe@41809: in berghofe@47438: set_env ctxt defs' types funs ids vcs' path prfx thy' berghofe@41809: end; berghofe@41809: berghofe@41809: end;