boehmes@36898: (* Title: HOL/Tools/SMT/smt_normalize.ML boehmes@36898: Author: Sascha Boehme, TU Muenchen boehmes@36898: boehmes@36898: Normalization steps on theorems required by SMT solvers: boehmes@36898: * simplify trivial distincts (those with less than three elements), boehmes@36898: * rewrite bool case expressions as if expressions, boehmes@36898: * normalize numerals (e.g. replace negative numerals by negated positive boehmes@36898: numerals), boehmes@36898: * embed natural numbers into integers, boehmes@36898: * add extra rules specifying types and constants which occur frequently, boehmes@36898: * fully translate into object logic, add universal closure, boehmes@39483: * monomorphize (create instances of schematic rules), boehmes@36898: * lift lambda terms, boehmes@36898: * make applications explicit for functions with varying number of arguments. boehmes@39483: * add (hypothetical definitions for) missing datatype selectors, boehmes@36898: *) boehmes@36898: boehmes@36898: signature SMT_NORMALIZE = boehmes@36898: sig boehmes@40162: type extra_norm = bool -> (int * thm) list -> Proof.context -> boehmes@40161: (int * thm) list * Proof.context boehmes@40424: val normalize: extra_norm -> bool -> (int * thm) list -> Proof.context -> boehmes@40161: (int * thm) list * Proof.context boehmes@36899: val atomize_conv: Proof.context -> conv boehmes@36898: val eta_expand_conv: (Proof.context -> conv) -> Proof.context -> conv boehmes@36898: end boehmes@36898: boehmes@36898: structure SMT_Normalize: SMT_NORMALIZE = boehmes@36898: struct boehmes@36898: boehmes@40663: structure U = SMT_Utils boehmes@40663: boehmes@36898: infix 2 ?? boehmes@36898: fun (test ?? f) x = if test x then f x else x boehmes@36898: boehmes@36898: boehmes@36898: boehmes@40685: (* instantiate elimination rules *) boehmes@40685: boehmes@40685: local boehmes@40685: val (cpfalse, cfalse) = `U.mk_cprop (Thm.cterm_of @{theory} @{const False}) boehmes@40685: boehmes@40685: fun inst f ct thm = boehmes@40685: let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm)) boehmes@40685: in Thm.instantiate ([], [(cv, ct)]) thm end boehmes@40685: in boehmes@40685: boehmes@40685: fun instantiate_elim thm = boehmes@40685: (case Thm.concl_of thm of boehmes@40685: @{const Trueprop} $ Var (_, @{typ bool}) => inst Thm.dest_arg cfalse thm boehmes@40685: | Var _ => inst I cpfalse thm boehmes@40685: | _ => thm) boehmes@40685: boehmes@40685: end boehmes@40685: boehmes@40685: boehmes@40685: boehmes@36898: (* simplification of trivial distincts (distinct should have at least boehmes@36898: three elements in the argument list) *) boehmes@36898: boehmes@36898: local boehmes@40681: fun is_trivial_distinct (Const (@{const_name distinct}, _) $ t) = boehmes@40681: (case try HOLogic.dest_list t of boehmes@40681: SOME [] => true boehmes@40681: | SOME [_] => true boehmes@40681: | _ => false) boehmes@36898: | is_trivial_distinct _ = false boehmes@36898: boehmes@37786: val thms = map mk_meta_eq @{lemma boehmes@40681: "distinct [] = True" boehmes@40681: "distinct [x] = True" boehmes@40681: "distinct [x, y] = (x ~= y)" boehmes@40681: by simp_all} boehmes@36898: fun distinct_conv _ = boehmes@40663: U.if_true_conv is_trivial_distinct (Conv.rewrs_conv thms) boehmes@36898: in boehmes@36898: fun trivial_distinct ctxt = boehmes@40161: map (apsnd ((Term.exists_subterm is_trivial_distinct o Thm.prop_of) ?? boehmes@40161: Conv.fconv_rule (Conv.top_conv distinct_conv ctxt))) boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: boehmes@36898: (* rewrite bool case expressions as if expressions *) boehmes@36898: boehmes@36898: local boehmes@36898: val is_bool_case = (fn boehmes@36898: Const (@{const_name "bool.bool_case"}, _) $ _ $ _ $ _ => true boehmes@36898: | _ => false) boehmes@36898: boehmes@40275: val thm = mk_meta_eq @{lemma boehmes@40275: "(case P of True => x | False => y) = (if P then x else y)" by simp} boehmes@40663: val unfold_conv = U.if_true_conv is_bool_case (Conv.rewr_conv thm) boehmes@36898: in boehmes@36898: fun rewrite_bool_cases ctxt = boehmes@40161: map (apsnd ((Term.exists_subterm is_bool_case o Thm.prop_of) ?? boehmes@40161: Conv.fconv_rule (Conv.top_conv (K unfold_conv) ctxt))) boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: boehmes@36898: (* normalization of numerals: rewriting of negative integer numerals into boehmes@36898: positive numerals, Numeral0 into 0, Numeral1 into 1 *) boehmes@36898: boehmes@36898: local boehmes@36898: fun is_number_sort ctxt T = boehmes@36898: Sign.of_sort (ProofContext.theory_of ctxt) (T, @{sort number_ring}) boehmes@36898: boehmes@36898: fun is_strange_number ctxt (t as Const (@{const_name number_of}, _) $ _) = boehmes@36898: (case try HOLogic.dest_number t of boehmes@36898: SOME (T, i) => is_number_sort ctxt T andalso i < 2 boehmes@36898: | NONE => false) boehmes@36898: | is_strange_number _ _ = false boehmes@36898: boehmes@36898: val pos_numeral_ss = HOL_ss boehmes@36898: addsimps [@{thm Int.number_of_minus}, @{thm Int.number_of_Min}] boehmes@36898: addsimps [@{thm Int.number_of_Pls}, @{thm Int.numeral_1_eq_1}] boehmes@36898: addsimps @{thms Int.pred_bin_simps} boehmes@36898: addsimps @{thms Int.normalize_bin_simps} boehmes@36898: addsimps @{lemma boehmes@36898: "Int.Min = - Int.Bit1 Int.Pls" boehmes@36898: "Int.Bit0 (- Int.Pls) = - Int.Pls" boehmes@36898: "Int.Bit0 (- k) = - Int.Bit0 k" boehmes@36898: "Int.Bit1 (- k) = - Int.Bit1 (Int.pred k)" boehmes@36898: by simp_all (simp add: pred_def)} boehmes@36898: boehmes@40663: fun pos_conv ctxt = U.if_conv (is_strange_number ctxt) boehmes@36898: (Simplifier.rewrite (Simplifier.context ctxt pos_numeral_ss)) boehmes@36898: Conv.no_conv boehmes@36898: in boehmes@36898: fun normalize_numerals ctxt = boehmes@40161: map (apsnd ((Term.exists_subterm (is_strange_number ctxt) o Thm.prop_of) ?? boehmes@40161: Conv.fconv_rule (Conv.top_sweep_conv pos_conv ctxt))) boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: boehmes@36898: (* embedding of standard natural number operations into integer operations *) boehmes@36898: boehmes@36898: local boehmes@40161: val nat_embedding = map (pair ~1) @{lemma boehmes@36898: "nat (int n) = n" boehmes@36898: "i >= 0 --> int (nat i) = i" boehmes@36898: "i < 0 --> int (nat i) = 0" boehmes@36898: by simp_all} boehmes@36898: boehmes@36898: val nat_rewriting = @{lemma boehmes@36898: "0 = nat 0" boehmes@36898: "1 = nat 1" boehmes@40279: "(number_of :: int => nat) = (%i. nat (number_of i))" boehmes@36898: "int (nat 0) = 0" boehmes@36898: "int (nat 1) = 1" boehmes@40279: "op < = (%a b. int a < int b)" boehmes@40279: "op <= = (%a b. int a <= int b)" boehmes@40279: "Suc = (%a. nat (int a + 1))" boehmes@40279: "op + = (%a b. nat (int a + int b))" boehmes@40279: "op - = (%a b. nat (int a - int b))" boehmes@40279: "op * = (%a b. nat (int a * int b))" boehmes@40279: "op div = (%a b. nat (int a div int b))" boehmes@40279: "op mod = (%a b. nat (int a mod int b))" boehmes@40279: "min = (%a b. nat (min (int a) (int b)))" boehmes@40279: "max = (%a b. nat (max (int a) (int b)))" boehmes@36898: "int (nat (int a + int b)) = int a + int b" boehmes@40279: "int (nat (int a + 1)) = int a + 1" (* special rule due to Suc above *) boehmes@36898: "int (nat (int a * int b)) = int a * int b" boehmes@36898: "int (nat (int a div int b)) = int a div int b" boehmes@36898: "int (nat (int a mod int b)) = int a mod int b" boehmes@36898: "int (nat (min (int a) (int b))) = min (int a) (int b)" boehmes@36898: "int (nat (max (int a) (int b))) = max (int a) (int b)" boehmes@40279: by (auto intro!: ext simp add: nat_mult_distrib nat_div_distrib boehmes@40279: nat_mod_distrib int_mult[symmetric] zdiv_int[symmetric] boehmes@40279: zmod_int[symmetric])} boehmes@36898: boehmes@36898: fun on_positive num f x = boehmes@36898: (case try HOLogic.dest_number (Thm.term_of num) of boehmes@36898: SOME (_, i) => if i >= 0 then SOME (f x) else NONE boehmes@36898: | NONE => NONE) boehmes@36898: boehmes@36898: val cancel_int_nat_ss = HOL_ss boehmes@36898: addsimps [@{thm Nat_Numeral.nat_number_of}] boehmes@36898: addsimps [@{thm Nat_Numeral.int_nat_number_of}] boehmes@36898: addsimps @{thms neg_simps} boehmes@36898: boehmes@40579: val int_eq = Thm.cterm_of @{theory} @{const "==" (int)} boehmes@40579: boehmes@36898: fun cancel_int_nat_simproc _ ss ct = boehmes@36898: let boehmes@36898: val num = Thm.dest_arg (Thm.dest_arg ct) boehmes@40579: val goal = Thm.mk_binop int_eq ct num boehmes@36898: val simpset = Simplifier.inherit_context ss cancel_int_nat_ss boehmes@36898: fun tac _ = Simplifier.simp_tac simpset 1 boehmes@36898: in on_positive num (Goal.prove_internal [] goal) tac end boehmes@36898: boehmes@36898: val nat_ss = HOL_ss boehmes@36898: addsimps nat_rewriting boehmes@40279: addsimprocs [ boehmes@40279: Simplifier.make_simproc { boehmes@40279: name = "cancel_int_nat_num", lhss = [@{cpat "int (nat _)"}], boehmes@40279: proc = cancel_int_nat_simproc, identifier = [] }] boehmes@36898: boehmes@36898: fun conv ctxt = Simplifier.rewrite (Simplifier.context ctxt nat_ss) boehmes@36898: boehmes@36898: val uses_nat_type = Term.exists_type (Term.exists_subtype (equal @{typ nat})) boehmes@40579: val uses_nat_int = Term.exists_subterm (member (op aconv) boehmes@40579: [@{const of_nat (int)}, @{const nat}]) boehmes@36898: in boehmes@36898: fun nat_as_int ctxt = boehmes@40161: map (apsnd ((uses_nat_type o Thm.prop_of) ?? Conv.fconv_rule (conv ctxt))) #> boehmes@40161: exists (uses_nat_int o Thm.prop_of o snd) ?? append nat_embedding boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: boehmes@36898: (* further normalizations: beta/eta, universal closure, atomize *) boehmes@36898: boehmes@36898: val eta_expand_eq = @{lemma "f == (%x. f x)" by (rule reflexive)} boehmes@36898: boehmes@36898: fun eta_expand_conv cv ctxt = boehmes@36898: Conv.rewr_conv eta_expand_eq then_conv Conv.abs_conv (cv o snd) ctxt boehmes@36898: boehmes@36898: local boehmes@36898: val eta_conv = eta_expand_conv boehmes@36898: boehmes@40279: fun args_conv cv ct = boehmes@40279: (case Thm.term_of ct of boehmes@40279: _ $ _ => Conv.combination_conv (args_conv cv) cv boehmes@40279: | _ => Conv.all_conv) ct boehmes@40279: boehmes@40279: fun eta_args_conv cv 0 = args_conv o cv boehmes@40279: | eta_args_conv cv i = eta_conv (eta_args_conv cv (i-1)) boehmes@40279: wenzelm@36936: fun keep_conv ctxt = Conv.binder_conv (norm_conv o snd) ctxt boehmes@36898: and eta_binder_conv ctxt = Conv.arg_conv (eta_conv norm_conv ctxt) boehmes@36898: and keep_let_conv ctxt = Conv.combination_conv boehmes@36898: (Conv.arg_conv (norm_conv ctxt)) (Conv.abs_conv (norm_conv o snd) ctxt) boehmes@36898: and unfold_let_conv ctxt = Conv.combination_conv boehmes@36898: (Conv.arg_conv (norm_conv ctxt)) (eta_conv norm_conv ctxt) boehmes@36898: and unfold_conv thm ctxt = Conv.rewr_conv thm then_conv keep_conv ctxt boehmes@36898: and unfold_ex1_conv ctxt = unfold_conv @{thm Ex1_def} ctxt boehmes@37786: and unfold_ball_conv ctxt = unfold_conv (mk_meta_eq @{thm Ball_def}) ctxt boehmes@37786: and unfold_bex_conv ctxt = unfold_conv (mk_meta_eq @{thm Bex_def}) ctxt boehmes@36898: and norm_conv ctxt ct = boehmes@36898: (case Thm.term_of ct of boehmes@36898: Const (@{const_name All}, _) $ Abs _ => keep_conv boehmes@36898: | Const (@{const_name All}, _) $ _ => eta_binder_conv boehmes@36898: | Const (@{const_name All}, _) => eta_conv eta_binder_conv boehmes@36898: | Const (@{const_name Ex}, _) $ Abs _ => keep_conv boehmes@36898: | Const (@{const_name Ex}, _) $ _ => eta_binder_conv boehmes@36898: | Const (@{const_name Ex}, _) => eta_conv eta_binder_conv boehmes@36898: | Const (@{const_name Let}, _) $ _ $ Abs _ => keep_let_conv boehmes@36898: | Const (@{const_name Let}, _) $ _ $ _ => unfold_let_conv boehmes@36898: | Const (@{const_name Let}, _) $ _ => eta_conv unfold_let_conv boehmes@36898: | Const (@{const_name Let}, _) => eta_conv (eta_conv unfold_let_conv) boehmes@36898: | Const (@{const_name Ex1}, _) $ _ => unfold_ex1_conv boehmes@36898: | Const (@{const_name Ex1}, _) => eta_conv unfold_ex1_conv boehmes@36898: | Const (@{const_name Ball}, _) $ _ $ _ => unfold_ball_conv boehmes@36898: | Const (@{const_name Ball}, _) $ _ => eta_conv unfold_ball_conv boehmes@36898: | Const (@{const_name Ball}, _) => eta_conv (eta_conv unfold_ball_conv) boehmes@36898: | Const (@{const_name Bex}, _) $ _ $ _ => unfold_bex_conv boehmes@36898: | Const (@{const_name Bex}, _) $ _ => eta_conv unfold_bex_conv boehmes@36898: | Const (@{const_name Bex}, _) => eta_conv (eta_conv unfold_bex_conv) boehmes@36898: | Abs _ => Conv.abs_conv (norm_conv o snd) boehmes@40279: | _ => boehmes@40279: (case Term.strip_comb (Thm.term_of ct) of boehmes@40279: (Const (c as (_, T)), ts) => boehmes@40279: if SMT_Builtin.is_builtin ctxt c boehmes@40279: then eta_args_conv norm_conv boehmes@40279: (length (Term.binder_types T) - length ts) boehmes@40279: else args_conv o norm_conv boehmes@40424: | _ => args_conv o norm_conv)) ctxt ct boehmes@36898: boehmes@40279: fun is_normed ctxt t = boehmes@36898: (case t of boehmes@40279: Const (@{const_name All}, _) $ Abs (_, _, u) => is_normed ctxt u boehmes@36898: | Const (@{const_name All}, _) $ _ => false boehmes@36898: | Const (@{const_name All}, _) => false boehmes@40279: | Const (@{const_name Ex}, _) $ Abs (_, _, u) => is_normed ctxt u boehmes@36898: | Const (@{const_name Ex}, _) $ _ => false boehmes@36898: | Const (@{const_name Ex}, _) => false boehmes@36898: | Const (@{const_name Let}, _) $ u1 $ Abs (_, _, u2) => boehmes@40279: is_normed ctxt u1 andalso is_normed ctxt u2 boehmes@36898: | Const (@{const_name Let}, _) $ _ $ _ => false boehmes@36898: | Const (@{const_name Let}, _) $ _ => false boehmes@36898: | Const (@{const_name Let}, _) => false boehmes@40279: | Const (@{const_name Ex1}, _) $ _ => false boehmes@36898: | Const (@{const_name Ex1}, _) => false boehmes@40279: | Const (@{const_name Ball}, _) $ _ $ _ => false boehmes@40279: | Const (@{const_name Ball}, _) $ _ => false boehmes@36898: | Const (@{const_name Ball}, _) => false boehmes@40279: | Const (@{const_name Bex}, _) $ _ $ _ => false boehmes@40279: | Const (@{const_name Bex}, _) $ _ => false boehmes@36898: | Const (@{const_name Bex}, _) => false boehmes@40279: | Abs (_, _, u) => is_normed ctxt u boehmes@40279: | _ => boehmes@40279: (case Term.strip_comb t of boehmes@40279: (Const (c as (_, T)), ts) => boehmes@40279: if SMT_Builtin.is_builtin ctxt c boehmes@40279: then length (Term.binder_types T) = length ts andalso boehmes@40279: forall (is_normed ctxt) ts boehmes@40279: else forall (is_normed ctxt) ts boehmes@40279: | (_, ts) => forall (is_normed ctxt) ts)) boehmes@36898: in boehmes@40279: fun norm_binder_conv ctxt = boehmes@40663: U.if_conv (is_normed ctxt) Conv.all_conv (norm_conv ctxt) boehmes@36898: end boehmes@36898: boehmes@36898: fun norm_def ctxt thm = boehmes@36898: (case Thm.prop_of thm of boehmes@40579: @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ _ $ Abs _) => boehmes@36898: norm_def ctxt (thm RS @{thm fun_cong}) boehmes@36898: | Const (@{const_name "=="}, _) $ _ $ Abs _ => boehmes@36898: norm_def ctxt (thm RS @{thm meta_eq_to_obj_eq}) boehmes@36898: | _ => thm) boehmes@36898: boehmes@36898: fun atomize_conv ctxt ct = boehmes@36898: (case Thm.term_of ct of boehmes@40579: @{const "==>"} $ _ $ _ => boehmes@36898: Conv.binop_conv (atomize_conv ctxt) then_conv boehmes@36898: Conv.rewr_conv @{thm atomize_imp} boehmes@36898: | Const (@{const_name "=="}, _) $ _ $ _ => boehmes@36898: Conv.binop_conv (atomize_conv ctxt) then_conv boehmes@36898: Conv.rewr_conv @{thm atomize_eq} boehmes@36898: | Const (@{const_name all}, _) $ Abs _ => wenzelm@36936: Conv.binder_conv (atomize_conv o snd) ctxt then_conv boehmes@36898: Conv.rewr_conv @{thm atomize_all} boehmes@36898: | _ => Conv.all_conv) ct boehmes@36898: boehmes@36898: fun normalize_rule ctxt = boehmes@36898: Conv.fconv_rule ( boehmes@36898: (* reduce lambda abstractions, except at known binders: *) boehmes@36898: Thm.beta_conversion true then_conv boehmes@36898: Thm.eta_conversion then_conv boehmes@36898: norm_binder_conv ctxt) #> boehmes@36898: norm_def ctxt #> boehmes@36898: Drule.forall_intr_vars #> boehmes@36898: Conv.fconv_rule (atomize_conv ctxt) boehmes@36898: boehmes@36898: boehmes@36898: boehmes@36898: (* lift lambda terms into additional rules *) boehmes@36898: boehmes@36898: local boehmes@36898: fun used_vars cvs ct = boehmes@36898: let boehmes@36898: val lookup = AList.lookup (op aconv) (map (` Thm.term_of) cvs) boehmes@36898: val add = (fn SOME ct => insert (op aconvc) ct | _ => I) boehmes@36898: in Term.fold_aterms (add o lookup) (Thm.term_of ct) [] end boehmes@36898: boehmes@36898: fun apply cv thm = boehmes@36898: let val thm' = Thm.combination thm (Thm.reflexive cv) boehmes@36898: in Thm.transitive thm' (Thm.beta_conversion false (Thm.rhs_of thm')) end boehmes@36898: fun apply_def cvs eq = Thm.symmetric (fold apply cvs eq) boehmes@36898: boehmes@36898: fun replace_lambda cvs ct (cx as (ctxt, defs)) = boehmes@36898: let boehmes@36898: val cvs' = used_vars cvs ct boehmes@36898: val ct' = fold_rev Thm.cabs cvs' ct boehmes@36898: in boehmes@36898: (case Termtab.lookup defs (Thm.term_of ct') of boehmes@36898: SOME eq => (apply_def cvs' eq, cx) boehmes@36898: | NONE => boehmes@36898: let boehmes@36898: val {T, ...} = Thm.rep_cterm ct' and n = Name.uu boehmes@36898: val (n', ctxt') = yield_singleton Variable.variant_fixes n ctxt boehmes@40663: val cu = U.mk_cequals (U.certify ctxt (Free (n', T))) ct' boehmes@36898: val (eq, ctxt'') = yield_singleton Assumption.add_assumes cu ctxt' boehmes@36898: val defs' = Termtab.update (Thm.term_of ct', eq) defs boehmes@36898: in (apply_def cvs' eq, (ctxt'', defs')) end) boehmes@36898: end boehmes@36898: boehmes@36898: fun none ct cx = (Thm.reflexive ct, cx) boehmes@36898: fun in_comb f g ct cx = boehmes@36898: let val (cu1, cu2) = Thm.dest_comb ct boehmes@36898: in cx |> f cu1 ||>> g cu2 |>> uncurry Thm.combination end boehmes@36898: fun in_arg f = in_comb none f boehmes@36898: fun in_abs f cvs ct (ctxt, defs) = boehmes@36898: let boehmes@36898: val (n, ctxt') = yield_singleton Variable.variant_fixes Name.uu ctxt boehmes@36898: val (cv, cu) = Thm.dest_abs (SOME n) ct boehmes@36898: in (ctxt', defs) |> f (cv :: cvs) cu |>> Thm.abstract_rule n cv end boehmes@36898: boehmes@36898: fun traverse cvs ct = boehmes@36898: (case Thm.term_of ct of boehmes@36898: Const (@{const_name All}, _) $ Abs _ => in_arg (in_abs traverse cvs) boehmes@36898: | Const (@{const_name Ex}, _) $ Abs _ => in_arg (in_abs traverse cvs) boehmes@36898: | Const (@{const_name Let}, _) $ _ $ Abs _ => boehmes@36898: in_comb (in_arg (traverse cvs)) (in_abs traverse cvs) boehmes@36898: | Abs _ => at_lambda cvs boehmes@36898: | _ $ _ => in_comb (traverse cvs) (traverse cvs) boehmes@36898: | _ => none) ct boehmes@36898: boehmes@36898: and at_lambda cvs ct = boehmes@36898: in_abs traverse cvs ct #-> (fn thm => boehmes@36898: replace_lambda cvs (Thm.rhs_of thm) #>> Thm.transitive thm) boehmes@36898: boehmes@36898: fun has_free_lambdas t = boehmes@36898: (case t of boehmes@36898: Const (@{const_name All}, _) $ Abs (_, _, u) => has_free_lambdas u boehmes@36898: | Const (@{const_name Ex}, _) $ Abs (_, _, u) => has_free_lambdas u boehmes@36898: | Const (@{const_name Let}, _) $ u1 $ Abs (_, _, u2) => boehmes@36898: has_free_lambdas u1 orelse has_free_lambdas u2 boehmes@36898: | Abs _ => true boehmes@36898: | u1 $ u2 => has_free_lambdas u1 orelse has_free_lambdas u2 boehmes@36898: | _ => false) boehmes@36898: boehmes@36898: fun lift_lm f thm cx = boehmes@36898: if not (has_free_lambdas (Thm.prop_of thm)) then (thm, cx) boehmes@36898: else cx |> f (Thm.cprop_of thm) |>> (fn thm' => Thm.equal_elim thm' thm) boehmes@36898: in boehmes@40161: fun lift_lambdas irules ctxt = boehmes@36898: let boehmes@36898: val cx = (ctxt, Termtab.empty) boehmes@40161: val (idxs, thms) = split_list irules boehmes@36898: val (thms', (ctxt', defs)) = fold_map (lift_lm (traverse [])) thms cx boehmes@36898: val eqs = Termtab.fold (cons o normalize_rule ctxt' o snd) defs [] boehmes@40161: in (map (pair ~1) eqs @ (idxs ~~ thms'), ctxt') end boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: boehmes@36898: (* make application explicit for functions with varying number of arguments *) boehmes@36898: boehmes@36898: local boehmes@36898: val const = prefix "c" and free = prefix "f" boehmes@36898: fun min i (e as (_, j)) = if i <> j then (true, Int.min (i, j)) else e boehmes@36898: fun add t i = Symtab.map_default (t, (false, i)) (min i) boehmes@36898: fun traverse t = boehmes@36898: (case Term.strip_comb t of boehmes@36898: (Const (n, _), ts) => add (const n) (length ts) #> fold traverse ts boehmes@36898: | (Free (n, _), ts) => add (free n) (length ts) #> fold traverse ts boehmes@36898: | (Abs (_, _, u), ts) => fold traverse (u :: ts) boehmes@36898: | (_, ts) => fold traverse ts) boehmes@40161: fun prune tab = Symtab.fold (fn (n, (true, i)) => boehmes@40161: Symtab.update (n, i) | _ => I) tab Symtab.empty boehmes@36898: boehmes@36898: fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 boehmes@36898: fun nary_conv conv1 conv2 ct = boehmes@36898: (Conv.combination_conv (nary_conv conv1 conv2) conv2 else_conv conv1) ct boehmes@36898: fun abs_conv conv tb = Conv.abs_conv (fn (cv, cx) => boehmes@36898: let val n = fst (Term.dest_Free (Thm.term_of cv)) boehmes@36898: in conv (Symtab.update (free n, 0) tb) cx end) boehmes@37153: val fun_app_rule = @{lemma "f x == fun_app f x" by (simp add: fun_app_def)} boehmes@36898: in boehmes@40161: fun explicit_application ctxt irules = boehmes@36898: let boehmes@36898: fun sub_conv tb ctxt ct = boehmes@36898: (case Term.strip_comb (Thm.term_of ct) of boehmes@36898: (Const (n, _), ts) => app_conv tb (const n) (length ts) ctxt boehmes@36898: | (Free (n, _), ts) => app_conv tb (free n) (length ts) ctxt boehmes@36898: | (Abs _, _) => nary_conv (abs_conv sub_conv tb ctxt) (sub_conv tb ctxt) boehmes@36898: | (_, _) => nary_conv Conv.all_conv (sub_conv tb ctxt)) ct boehmes@36898: and app_conv tb n i ctxt = boehmes@36898: (case Symtab.lookup tb n of boehmes@36898: NONE => nary_conv Conv.all_conv (sub_conv tb ctxt) boehmes@37153: | SOME j => fun_app_conv tb ctxt (i - j)) boehmes@37153: and fun_app_conv tb ctxt i ct = ( boehmes@36898: if i = 0 then nary_conv Conv.all_conv (sub_conv tb ctxt) boehmes@36898: else boehmes@37153: Conv.rewr_conv fun_app_rule then_conv boehmes@37153: binop_conv (fun_app_conv tb ctxt (i-1)) (sub_conv tb ctxt)) ct boehmes@36898: boehmes@36898: fun needs_exp_app tab = Term.exists_subterm (fn boehmes@36898: Bound _ $ _ => true boehmes@36898: | Const (n, _) => Symtab.defined tab (const n) boehmes@36898: | Free (n, _) => Symtab.defined tab (free n) boehmes@36898: | _ => false) boehmes@36898: boehmes@36898: fun rewrite tab ctxt thm = boehmes@36898: if not (needs_exp_app tab (Thm.prop_of thm)) then thm boehmes@36898: else Conv.fconv_rule (sub_conv tab ctxt) thm boehmes@36898: boehmes@40161: val tab = prune (fold (traverse o Thm.prop_of o snd) irules Symtab.empty) boehmes@40161: in map (apsnd (rewrite tab ctxt)) irules end boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: boehmes@39483: (* add missing datatype selectors via hypothetical definitions *) boehmes@39483: boehmes@39483: local boehmes@39483: val add = (fn Type (n, _) => Symtab.update (n, ()) | _ => I) boehmes@39483: boehmes@39483: fun collect t = boehmes@39483: (case Term.strip_comb t of boehmes@39483: (Abs (_, T, t), _) => add T #> collect t boehmes@39483: | (Const (_, T), ts) => collects T ts boehmes@39483: | (Free (_, T), ts) => collects T ts boehmes@39483: | _ => I) boehmes@39483: and collects T ts = boehmes@39483: let val ((Ts, Us), U) = Term.strip_type T |> apfst (chop (length ts)) boehmes@39483: in fold add Ts #> add (Us ---> U) #> fold collect ts end boehmes@39483: boehmes@39483: fun add_constructors thy n = boehmes@39483: (case Datatype.get_info thy n of boehmes@39483: NONE => I boehmes@39483: | SOME {descr, ...} => fold (fn (_, (_, _, cs)) => fold (fn (n, ds) => boehmes@39483: fold (insert (op =) o pair n) (1 upto length ds)) cs) descr) boehmes@39483: boehmes@39483: fun add_selector (c as (n, i)) ctxt = boehmes@39483: (case Datatype_Selectors.lookup_selector ctxt c of boehmes@39483: SOME _ => ctxt boehmes@39483: | NONE => boehmes@39483: let boehmes@39483: val T = Sign.the_const_type (ProofContext.theory_of ctxt) n boehmes@39483: val U = Term.body_type T --> nth (Term.binder_types T) (i-1) boehmes@39483: in boehmes@39483: ctxt boehmes@39483: |> yield_singleton Variable.variant_fixes Name.uu boehmes@39483: |>> pair ((n, T), i) o rpair U boehmes@39483: |-> Context.proof_map o Datatype_Selectors.add_selector boehmes@39483: end) boehmes@39483: in boehmes@39483: boehmes@40161: fun datatype_selectors irules ctxt = boehmes@39483: let boehmes@40161: val ns = Symtab.keys (fold (collect o Thm.prop_of o snd) irules Symtab.empty) boehmes@39483: val cs = fold (add_constructors (ProofContext.theory_of ctxt)) ns [] boehmes@40161: in (irules, fold add_selector cs ctxt) end boehmes@39483: (* FIXME: also generate hypothetical definitions for the selectors *) boehmes@39483: boehmes@39483: end boehmes@39483: boehmes@39483: boehmes@39483: boehmes@36898: (* combined normalization *) boehmes@36898: boehmes@40162: type extra_norm = bool -> (int * thm) list -> Proof.context -> boehmes@40161: (int * thm) list * Proof.context boehmes@36898: boehmes@40161: fun with_context f irules ctxt = (f ctxt irules, ctxt) boehmes@36898: boehmes@40424: fun normalize extra_norm with_datatypes irules ctxt = boehmes@40278: let boehmes@40278: fun norm f ctxt' (i, thm) = boehmes@40424: if Config.get ctxt' SMT_Config.drop_bad_facts then boehmes@40278: (case try (f ctxt') thm of boehmes@40278: SOME thm' => SOME (i, thm') boehmes@40424: | NONE => (SMT_Config.verbose_msg ctxt' (prefix ("Warning: " ^ boehmes@40278: "dropping assumption: ") o Display.string_of_thm ctxt') thm; NONE)) boehmes@40424: else SOME (i, f ctxt' thm) boehmes@40278: in boehmes@40278: irules boehmes@40685: |> map (apsnd instantiate_elim) boehmes@40278: |> trivial_distinct ctxt boehmes@40278: |> rewrite_bool_cases ctxt boehmes@40278: |> normalize_numerals ctxt boehmes@40278: |> nat_as_int ctxt boehmes@40278: |> rpair ctxt boehmes@40278: |-> extra_norm with_datatypes boehmes@40278: |-> with_context (map_filter o norm normalize_rule) boehmes@40278: |-> SMT_Monomorph.monomorph boehmes@40278: |-> lift_lambdas boehmes@40278: |-> with_context explicit_application boehmes@40278: |-> (if with_datatypes then datatype_selectors else pair) boehmes@40278: end boehmes@36898: boehmes@36898: end