haftmann@24219: (* Title: Pure/Isar/code_unit.ML haftmann@24219: Author: Florian Haftmann, TU Muenchen haftmann@24219: haftmann@24844: Basic notions of code generation. Auxiliary. haftmann@24219: *) haftmann@24219: haftmann@24219: signature CODE_UNIT = haftmann@24219: sig haftmann@24844: (*generic non-sense*) haftmann@24844: val bad_thm: string -> 'a haftmann@28368: val error_thm: ('a -> 'b) -> 'a -> 'b haftmann@28368: val warning_thm: ('a -> 'b) -> 'a -> 'b option haftmann@28368: val try_thm: ('a -> 'b) -> 'a -> 'b option haftmann@24844: haftmann@24844: (*typ instantiations*) haftmann@26970: val typscheme: theory -> string * typ -> (string * sort) list * typ haftmann@28423: val inst_thm: theory -> sort Vartab.table -> thm -> thm haftmann@28423: val constrain_thm: theory -> sort -> thm -> thm haftmann@24844: haftmann@26747: (*constant aliasses*) haftmann@26354: val add_const_alias: thm -> theory -> theory haftmann@28346: val triv_classes: theory -> class list haftmann@26747: val resubst_alias: theory -> string -> string haftmann@26747: haftmann@26747: (*constants*) haftmann@24219: val string_of_typ: theory -> typ -> string haftmann@24423: val string_of_const: theory -> string -> string haftmann@24423: val no_args: theory -> string -> int haftmann@26112: val check_const: theory -> term -> string haftmann@24219: val read_bare_const: theory -> string -> string * typ haftmann@24423: val read_const: theory -> string -> string haftmann@24219: haftmann@24844: (*constructor sets*) haftmann@24423: val constrset_of_consts: theory -> (string * typ) list haftmann@24219: -> string * ((string * sort) list * (string * typ list) list) haftmann@24219: haftmann@30022: (*code equations*) haftmann@28423: val assert_eqn: theory -> thm -> thm haftmann@28423: val mk_eqn: theory -> thm -> thm * bool haftmann@28708: val assert_linear: (string -> bool) -> thm * bool -> thm * bool haftmann@28423: val const_eqn: thm -> string haftmann@28423: val const_typ_eqn: thm -> string * typ haftmann@28423: val head_eqn: theory -> thm -> string * ((string * sort) list * typ) haftmann@28423: val expand_eta: theory -> int -> thm -> thm haftmann@28368: val rewrite_eqn: simpset -> thm -> thm haftmann@27582: val rewrite_head: thm list -> thm -> thm haftmann@28423: val norm_args: theory -> thm list -> thm list haftmann@28423: val norm_varnames: theory -> (string -> string) -> (string -> string) -> thm list -> thm list haftmann@24844: haftmann@24844: (*case certificates*) haftmann@24844: val case_cert: thm -> string * (int * string list) haftmann@24219: end; haftmann@24219: haftmann@28054: structure Code_Unit: CODE_UNIT = haftmann@24219: struct haftmann@24219: haftmann@24219: haftmann@24219: (* auxiliary *) haftmann@24219: haftmann@24219: exception BAD_THM of string; haftmann@24219: fun bad_thm msg = raise BAD_THM msg; haftmann@24219: fun error_thm f thm = f thm handle BAD_THM msg => error msg; haftmann@24219: fun warning_thm f thm = SOME (f thm) handle BAD_THM msg haftmann@24219: => (warning ("code generator: " ^ msg); NONE); haftmann@24624: fun try_thm f thm = SOME (f thm) handle BAD_THM _ => NONE; haftmann@24219: wenzelm@26939: fun string_of_typ thy = setmp show_sorts true (Syntax.string_of_typ_global thy); haftmann@25597: fun string_of_const thy c = case AxClass.inst_of_param thy c haftmann@24423: of SOME (c, tyco) => Sign.extern_const thy c ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco) haftmann@24423: | NONE => Sign.extern_const thy c; haftmann@24219: haftmann@24423: fun no_args thy = length o fst o strip_type o Sign.the_const_type thy; haftmann@24219: haftmann@26354: haftmann@26354: (* utilities *) haftmann@26354: haftmann@26970: fun typscheme thy (c, ty) = haftmann@26970: let haftmann@30022: val ty' = Logic.unvarifyT ty; haftmann@30022: fun dest (TFree (v, sort)) = (v, sort) haftmann@26970: | dest ty = error ("Illegal type parameter in type scheme: " ^ Syntax.string_of_typ_global thy ty); haftmann@30022: val vs = map dest (Sign.const_typargs thy (c, ty')); haftmann@30022: in (vs, Type.strip_sorts ty') end; haftmann@26970: haftmann@28423: fun inst_thm thy tvars' thm = haftmann@26354: let haftmann@26354: val tvars = (Term.add_tvars o Thm.prop_of) thm []; haftmann@28310: val inter_sort = Sorts.inter_sort (Sign.classes_of thy); haftmann@28310: fun mk_inst (tvar as (v, sort)) = case Vartab.lookup tvars' v haftmann@28310: of SOME sort' => SOME (pairself (Thm.ctyp_of thy o TVar) haftmann@28310: (tvar, (v, inter_sort (sort, sort')))) haftmann@26354: | NONE => NONE; haftmann@26354: val insts = map_filter mk_inst tvars; haftmann@26354: in Thm.instantiate (insts, []) thm end; haftmann@26354: haftmann@28423: fun constrain_thm thy sort thm = haftmann@26354: let haftmann@26354: val constrain = curry (Sorts.inter_sort (Sign.classes_of thy)) sort haftmann@26354: val tvars = (Term.add_tvars o Thm.prop_of) thm []; haftmann@26354: fun mk_inst (tvar as (v, sort)) = pairself (Thm.ctyp_of thy o TVar o pair v) haftmann@26354: (sort, constrain sort) haftmann@26354: val insts = map mk_inst tvars; haftmann@26354: in Thm.instantiate (insts, []) thm end; haftmann@26354: haftmann@28423: fun expand_eta thy k thm = haftmann@26354: let haftmann@26354: val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm; haftmann@26354: val (head, args) = strip_comb lhs; haftmann@26354: val l = if k = ~1 haftmann@26354: then (length o fst o strip_abs) rhs haftmann@26354: else Int.max (0, k - length args); haftmann@26354: val used = Name.make_context (map (fst o fst) (Term.add_vars lhs [])); haftmann@26354: fun get_name _ 0 = pair [] haftmann@26354: | get_name (Abs (v, ty, t)) k = haftmann@26354: Name.variants [v] haftmann@26354: ##>> get_name t (k - 1) haftmann@26354: #>> (fn ([v'], vs') => (v', ty) :: vs') haftmann@26354: | get_name t k = haftmann@26354: let haftmann@26354: val (tys, _) = (strip_type o fastype_of) t haftmann@26354: in case tys haftmann@26354: of [] => raise TERM ("expand_eta", [t]) haftmann@26354: | ty :: _ => haftmann@26354: Name.variants [""] haftmann@26354: #-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1) haftmann@26354: #>> (fn vs' => (v, ty) :: vs')) haftmann@26354: end; haftmann@26354: val (vs, _) = get_name rhs l used; haftmann@26354: fun expand (v, ty) thm = Drule.fun_cong_rule thm haftmann@26354: (Thm.cterm_of thy (Var ((v, 0), ty))); haftmann@26354: in haftmann@26354: thm haftmann@26354: |> fold expand vs haftmann@26354: |> Conv.fconv_rule Drule.beta_eta_conversion haftmann@26354: end; haftmann@26354: haftmann@28368: fun eqn_conv conv = haftmann@26354: let haftmann@26354: fun lhs_conv ct = if can Thm.dest_comb ct haftmann@26354: then (Conv.combination_conv lhs_conv conv) ct haftmann@26354: else Conv.all_conv ct; haftmann@26354: in Conv.combination_conv (Conv.arg_conv lhs_conv) conv end; haftmann@26354: haftmann@26354: fun head_conv conv = haftmann@26354: let haftmann@26354: fun lhs_conv ct = if can Thm.dest_comb ct haftmann@26354: then (Conv.fun_conv lhs_conv) ct haftmann@26354: else conv ct; haftmann@26354: in Conv.fun_conv (Conv.arg_conv lhs_conv) end; haftmann@26354: haftmann@28368: val rewrite_eqn = Conv.fconv_rule o eqn_conv o Simplifier.rewrite; haftmann@26354: val rewrite_head = Conv.fconv_rule o head_conv o MetaSimplifier.rewrite false; haftmann@26354: haftmann@28423: fun norm_args thy thms = haftmann@26354: let haftmann@26354: val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals; haftmann@28310: val k = fold (curry Int.max o num_args_of o Thm.prop_of) thms 0; haftmann@26354: in haftmann@26354: thms haftmann@28423: |> map (expand_eta thy k) haftmann@26354: |> map (Conv.fconv_rule Drule.beta_eta_conversion) haftmann@26354: end; haftmann@26354: haftmann@28423: fun canonical_tvars thy purify_tvar thm = haftmann@26354: let haftmann@28423: val ctyp = Thm.ctyp_of thy; haftmann@26354: fun tvars_subst_for thm = (fold_types o fold_atyps) haftmann@26354: (fn TVar (v_i as (v, _), sort) => let haftmann@26354: val v' = purify_tvar v haftmann@26354: in if v = v' then I haftmann@26354: else insert (op =) (v_i, (v', sort)) end haftmann@26354: | _ => I) (prop_of thm) []; haftmann@26354: fun mk_inst (v_i, (v', sort)) (maxidx, acc) = haftmann@26354: let haftmann@26354: val ty = TVar (v_i, sort) haftmann@26354: in haftmann@26354: (maxidx + 1, (ctyp ty, ctyp (TVar ((v', maxidx), sort))) :: acc) haftmann@26354: end; haftmann@26354: val maxidx = Thm.maxidx_of thm + 1; haftmann@26354: val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []); haftmann@26354: in Thm.instantiate (inst, []) thm end; haftmann@26354: haftmann@28423: fun canonical_vars thy purify_var thm = haftmann@26354: let haftmann@28423: val cterm = Thm.cterm_of thy; haftmann@26354: fun vars_subst_for thm = fold_aterms haftmann@26354: (fn Var (v_i as (v, _), ty) => let haftmann@26354: val v' = purify_var v haftmann@26354: in if v = v' then I haftmann@26354: else insert (op =) (v_i, (v', ty)) end haftmann@26354: | _ => I) (prop_of thm) []; haftmann@26354: fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) = haftmann@26354: let haftmann@26354: val t = Var (v_i, ty) haftmann@26354: in haftmann@26354: (maxidx + 1, (cterm t, cterm (Var ((v', maxidx), ty))) :: acc) haftmann@26354: end; haftmann@26354: val maxidx = Thm.maxidx_of thm + 1; haftmann@26354: val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []); haftmann@26354: in Thm.instantiate ([], inst) thm end; haftmann@26354: haftmann@26354: fun canonical_absvars purify_var thm = haftmann@26354: let haftmann@26354: val t = Thm.plain_prop_of thm; haftmann@26354: val t' = Term.map_abs_vars purify_var t; haftmann@26354: in Thm.rename_boundvars t t' thm end; haftmann@26354: haftmann@28423: fun norm_varnames thy purify_tvar purify_var thms = haftmann@26354: let haftmann@26354: fun burrow_thms f [] = [] haftmann@26354: | burrow_thms f thms = haftmann@26354: thms haftmann@26354: |> Conjunction.intr_balanced haftmann@26354: |> f haftmann@26354: |> Conjunction.elim_balanced (length thms) haftmann@26354: in haftmann@26354: thms haftmann@28423: |> burrow_thms (canonical_tvars thy purify_tvar) haftmann@28423: |> map (canonical_vars thy purify_var) haftmann@26354: |> map (canonical_absvars purify_var) haftmann@26354: |> map Drule.zero_var_indexes haftmann@26354: end; haftmann@26354: haftmann@28346: haftmann@26354: (* const aliasses *) haftmann@26354: haftmann@26354: structure ConstAlias = TheoryDataFun haftmann@26354: ( haftmann@26747: type T = ((string * string) * thm) list * class list; haftmann@26747: val empty = ([], []); haftmann@26354: val copy = I; wenzelm@26618: val extend = I; wenzelm@29288: fun merge _ ((alias1, classes1), (alias2, classes2)) : T = haftmann@26747: (Library.merge (eq_snd Thm.eq_thm_prop) (alias1, alias2), haftmann@26747: Library.merge (op =) (classes1, classes2)); haftmann@26354: ); haftmann@26354: haftmann@28423: fun add_const_alias thm thy = haftmann@26354: let haftmann@28423: val lhs_rhs = case try Logic.dest_equals (Thm.prop_of thm) haftmann@26354: of SOME lhs_rhs => lhs_rhs haftmann@26354: | _ => error ("Not an equation: " ^ Display.string_of_thm thm); haftmann@26354: val c_c' = case try (pairself (AxClass.unoverload_const thy o dest_Const)) lhs_rhs haftmann@26354: of SOME c_c' => c_c' haftmann@26354: | _ => error ("Not an equation with two constants: " ^ Display.string_of_thm thm); haftmann@26747: val some_class = the_list (AxClass.class_of_param thy (snd c_c')); haftmann@28423: in thy |> haftmann@26747: ConstAlias.map (fn (alias, classes) => haftmann@26747: ((c_c', thm) :: alias, fold (insert (op =)) some_class classes)) haftmann@26747: end; haftmann@26354: haftmann@26747: fun resubst_alias thy = haftmann@26747: let haftmann@26747: val alias = fst (ConstAlias.get thy); haftmann@26747: val subst_inst_param = Option.map fst o AxClass.inst_of_param thy; haftmann@26747: fun subst_alias c = haftmann@26747: get_first (fn ((c', c''), _) => if c = c'' then SOME c' else NONE) alias; haftmann@26747: in haftmann@26747: perhaps subst_inst_param haftmann@26747: #> perhaps subst_alias haftmann@26747: end; haftmann@26747: haftmann@26747: val triv_classes = snd o ConstAlias.get; haftmann@26747: haftmann@28346: wenzelm@26610: (* reading constants as terms *) haftmann@24219: haftmann@26112: fun check_bare_const thy t = case try dest_Const t haftmann@26112: of SOME c_ty => c_ty wenzelm@26939: | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t); haftmann@26112: haftmann@28346: fun check_const thy = AxClass.unoverload_const thy o check_bare_const thy; haftmann@26112: haftmann@26112: fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy; haftmann@24219: haftmann@28346: fun read_const thy = AxClass.unoverload_const thy o read_bare_const thy; haftmann@24219: haftmann@24219: haftmann@24423: (* constructor sets *) haftmann@24219: haftmann@24423: fun constrset_of_consts thy cs = haftmann@24219: let haftmann@28704: val _ = map (fn (c, _) => if (is_some o AxClass.class_of_param thy) c haftmann@28704: then error ("Is a class parameter: " ^ string_of_const thy c) else ()) cs; haftmann@24423: fun no_constr (c, ty) = error ("Not a datatype constructor: " ^ string_of_const thy c haftmann@24423: ^ " :: " ^ string_of_typ thy ty); haftmann@24423: fun last_typ c_ty ty = haftmann@24423: let wenzelm@29270: val frees = OldTerm.typ_tfrees ty; haftmann@24423: val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) ty haftmann@24423: handle TYPE _ => no_constr c_ty haftmann@24423: val _ = if has_duplicates (eq_fst (op =)) vs then no_constr c_ty else (); haftmann@24423: val _ = if length frees <> length vs then no_constr c_ty else (); haftmann@24423: in (tyco, vs) end; haftmann@24423: fun ty_sorts (c, ty) = haftmann@24423: let haftmann@24423: val ty_decl = (Logic.unvarifyT o Sign.the_const_type thy) c; haftmann@26970: val (tyco, _) = last_typ (c, ty) ty_decl; haftmann@24423: val (_, vs) = last_typ (c, ty) ty; haftmann@28015: in ((tyco, map snd vs), (c, (map fst vs, ty))) end; haftmann@24423: fun add ((tyco', sorts'), c) ((tyco, sorts), cs) = haftmann@24423: let haftmann@24423: val _ = if tyco' <> tyco haftmann@24423: then error "Different type constructors in constructor set" haftmann@24423: else (); haftmann@24423: val sorts'' = map2 (curry (Sorts.inter_sort (Sign.classes_of thy))) sorts' sorts haftmann@24423: in ((tyco, sorts), c :: cs) end; haftmann@24423: fun inst vs' (c, (vs, ty)) = haftmann@24423: let haftmann@24423: val the_v = the o AList.lookup (op =) (vs ~~ vs'); haftmann@24423: val ty' = map_atyps (fn TFree (v, _) => TFree (the_v v)) ty; haftmann@24423: in (c, (fst o strip_type) ty') end; haftmann@24423: val c' :: cs' = map ty_sorts cs; haftmann@24423: val ((tyco, sorts), cs'') = fold add cs' (apsnd single c'); wenzelm@24848: val vs = Name.names Name.context Name.aT sorts; haftmann@24423: val cs''' = map (inst vs) cs''; haftmann@30022: in (tyco, (vs, rev cs''')) end; haftmann@24219: haftmann@24219: haftmann@30022: (* code equations *) haftmann@24219: haftmann@28423: fun assert_eqn thy thm = haftmann@24219: let haftmann@24219: val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm haftmann@24219: handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm) haftmann@24219: | THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm); haftmann@24219: fun vars_of t = fold_aterms haftmann@24219: (fn Var (v, _) => insert (op =) v haftmann@24219: | Free _ => bad_thm ("Illegal free variable in rewrite theorem\n" haftmann@24219: ^ Display.string_of_thm thm) haftmann@24219: | _ => I) t []; haftmann@24219: fun tvars_of t = fold_term_types haftmann@24219: (fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v haftmann@24219: | TFree _ => bad_thm haftmann@24219: ("Illegal free type variable in rewrite theorem\n" ^ Display.string_of_thm thm))) t []; haftmann@24219: val lhs_vs = vars_of lhs; haftmann@24219: val rhs_vs = vars_of rhs; haftmann@24219: val lhs_tvs = tvars_of lhs; haftmann@24219: val rhs_tvs = tvars_of lhs; haftmann@24219: val _ = if null (subtract (op =) lhs_vs rhs_vs) haftmann@24219: then () haftmann@24219: else bad_thm ("Free variables on right hand side of rewrite theorem\n" haftmann@24219: ^ Display.string_of_thm thm); haftmann@24219: val _ = if null (subtract (op =) lhs_tvs rhs_tvs) haftmann@24219: then () haftmann@24219: else bad_thm ("Free type variables on right hand side of rewrite theorem\n" haftmann@28423: ^ Display.string_of_thm thm) val (head, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm; haftmann@28423: val (c, ty) = case head of Const c_ty => c_ty | _ => haftmann@24219: bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm thm); haftmann@24219: fun check _ (Abs _) = bad_thm haftmann@24219: ("Abstraction on left hand side of equation\n" haftmann@24219: ^ Display.string_of_thm thm) haftmann@24219: | check 0 (Var _) = () haftmann@24219: | check _ (Var _) = bad_thm haftmann@30022: ("Variable with application on left hand side of code equation\n" haftmann@24219: ^ Display.string_of_thm thm) haftmann@24219: | check n (t1 $ t2) = (check (n+1) t1; check 0 t2) haftmann@24219: | check n (Const (_, ty)) = if n <> (length o fst o strip_type) ty haftmann@24219: then bad_thm haftmann@24219: ("Partially applied constant on left hand side of equation\n" haftmann@24219: ^ Display.string_of_thm thm) haftmann@24219: else (); haftmann@24219: val _ = map (check 0) args; haftmann@28423: val ty_decl = Sign.the_const_type thy c; haftmann@28423: val _ = if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty) haftmann@28423: then () else bad_thm ("Type\n" ^ string_of_typ thy ty haftmann@30022: ^ "\nof code equation\n" haftmann@28423: ^ Display.string_of_thm thm haftmann@28423: ^ "\nis incompatible with declared function type\n" haftmann@28423: ^ string_of_typ thy ty_decl) haftmann@28423: in thm end; haftmann@24219: haftmann@28423: fun add_linear thm = haftmann@24219: let haftmann@28423: val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm; haftmann@28423: val linear = not (has_duplicates (op =) haftmann@28423: ((fold o fold_aterms) (fn Var (v, _) => cons v | _ => I) args [])) haftmann@28423: in (thm, linear) end; haftmann@28423: haftmann@28708: fun assert_pat is_cons thm = haftmann@28708: let haftmann@28708: val args = (snd o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm; haftmann@28708: val _ = (map o map_aterms) (fn t as Const (c, _) => if is_cons c then t haftmann@28708: else bad_thm ("Not a constructor on left hand side of equation: " haftmann@28708: ^ quote c ^ ",\n in equation\n" ^ Display.string_of_thm thm) haftmann@28708: | t => t) args; haftmann@28708: in thm end; haftmann@28708: haftmann@28708: fun assert_linear is_cons (thm, false) = (thm, false) haftmann@28708: | assert_linear is_cons (thm, true) = if snd (add_linear (assert_pat is_cons thm)) then (thm, true) haftmann@28423: else bad_thm haftmann@30022: ("Duplicate variables on left hand side of code equation:\n" haftmann@28423: ^ Display.string_of_thm thm); haftmann@28423: haftmann@28423: haftmann@28423: fun mk_eqn thy = add_linear o assert_eqn thy o AxClass.unoverload thy haftmann@28423: o LocalDefs.meta_rewrite_rule (ProofContext.init thy); haftmann@28423: haftmann@28423: val const_typ_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of; haftmann@28423: val const_eqn = fst o const_typ_eqn; haftmann@28423: fun head_eqn thy thm = let val (c, ty) = const_typ_eqn thm in (c, typscheme thy (c, ty)) end; haftmann@24219: haftmann@24219: haftmann@24844: (* case cerificates *) haftmann@24844: haftmann@24917: fun case_certificate thm = haftmann@24844: let haftmann@24917: val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals haftmann@24917: o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.prop_of) thm; haftmann@24844: val _ = case head of Free _ => true haftmann@24844: | Var _ => true haftmann@24844: | _ => raise TERM ("case_cert", []); haftmann@24917: val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr; haftmann@24917: val (Const (case_const, _), raw_params) = strip_comb case_expr; haftmann@24917: val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params; haftmann@24917: val _ = if n = ~1 then raise TERM ("case_cert", []) else (); haftmann@24917: val params = map (fst o dest_Var) (nth_drop n raw_params); haftmann@24917: fun dest_case t = haftmann@24844: let haftmann@24844: val (head' $ t_co, rhs) = Logic.dest_equals t; haftmann@24844: val _ = if head' = head then () else raise TERM ("case_cert", []); haftmann@24844: val (Const (co, _), args) = strip_comb t_co; haftmann@24844: val (Var (param, _), args') = strip_comb rhs; haftmann@24844: val _ = if args' = args then () else raise TERM ("case_cert", []); haftmann@24844: in (param, co) end; haftmann@24917: fun analyze_cases cases = haftmann@24844: let haftmann@24917: val co_list = fold (AList.update (op =) o dest_case) cases []; haftmann@24844: in map (the o AList.lookup (op =) co_list) params end; haftmann@24844: fun analyze_let t = haftmann@24844: let haftmann@24844: val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t; haftmann@24844: val _ = if head' = head then () else raise TERM ("case_cert", []); haftmann@24844: val _ = if arg' = arg then () else raise TERM ("case_cert", []); haftmann@24844: val _ = if [param'] = params then () else raise TERM ("case_cert", []); haftmann@24844: in [] end; haftmann@24917: fun analyze (cases as [let_case]) = haftmann@24917: (analyze_cases cases handle Bind => analyze_let let_case) haftmann@24917: | analyze cases = analyze_cases cases; haftmann@24917: in (case_const, (n, analyze cases)) end; haftmann@24917: haftmann@24917: fun case_cert thm = case_certificate thm haftmann@24917: handle Bind => error "bad case certificate" haftmann@28346: | TERM _ => error "bad case certificate"; haftmann@24844: haftmann@24219: end;