| author | paulson <lp15@cam.ac.uk> | 
| Sat, 02 Nov 2019 14:31:34 +0000 | |
| changeset 70999 | 5b753486c075 | 
| parent 70479 | 02d08d0ba896 | 
| child 80634 | a90ab1ea6458 | 
| permissions | -rw-r--r-- | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 1 | (* Title: HOL/Tools/Function/function_elims.ML | 
| 53609 | 2 | Author: Manuel Eberl, TU Muenchen | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 3 | |
| 53666 | 4 | Generate the pelims rules for a function. These are of the shape | 
| 53664 | 5 | [|f x y z = w; !!\<dots>. [|x = \<dots>; y = \<dots>; z = \<dots>; w = \<dots>|] ==> P; \<dots>|] ==> P | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 6 | and are derived from the cases rule. There is at least one pelim rule for | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 7 | each function (cf. mutually recursive functions) | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 8 | There may be more than one pelim rule for a function in case of functions | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 9 | that return a boolean. For such a function, e.g. P x, not only the normal | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 10 | elim rule with the premise P x = z is generated, but also two additional | 
| 53664 | 11 | elim rules with P x resp. \<not>P x as premises. | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 12 | *) | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 13 | |
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 14 | signature FUNCTION_ELIMS = | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 15 | sig | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 16 | val dest_funprop : term -> (term * term list) * term | 
| 53667 
0aefb31e27e0
distinguish Proof.context vs. local_theory semantically, with corresponding naming conventions;
 wenzelm parents: 
53666diff
changeset | 17 | val mk_partial_elim_rules : Proof.context -> | 
| 53666 | 18 | Function_Common.function_result -> thm list list | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 19 | end; | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 20 | |
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 21 | structure Function_Elims : FUNCTION_ELIMS = | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 22 | struct | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 23 | |
| 53669 
6ede465b5be8
more antiquotations -- avoid unchecked string literals;
 wenzelm parents: 
53667diff
changeset | 24 | (* Extract a function and its arguments from a proposition that is | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 25 | either of the form "f x y z = ..." or, in case of function that | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 26 | returns a boolean, "f x y z" *) | 
| 67149 | 27 | fun dest_funprop (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ lhs $ rhs) = (strip_comb lhs, rhs) | 
| 28 | | dest_funprop (Const (\<^const_name>\<open>Not\<close>, _) $ t) = (strip_comb t, \<^term>\<open>False\<close>) | |
| 29 | | dest_funprop t = (strip_comb t, \<^term>\<open>True\<close>); | |
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 30 | |
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 31 | local | 
| 53666 | 32 | |
| 54738 | 33 | fun propagate_tac ctxt i = | 
| 54736 | 34 | let | 
| 35 | fun inspect eq = | |
| 36 | (case eq of | |
| 67149 | 37 | Const (\<^const_name>\<open>Trueprop\<close>, _) $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ Free x $ t) => | 
| 54736 | 38 | if Logic.occs (Free x, t) then raise Match else true | 
| 67149 | 39 | | Const (\<^const_name>\<open>Trueprop\<close>, _) $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t $ Free x) => | 
| 54736 | 40 | if Logic.occs (Free x, t) then raise Match else false | 
| 41 | | _ => raise Match); | |
| 42 | fun mk_eq thm = | |
| 59582 | 43 | (if inspect (Thm.prop_of thm) then [thm RS eq_reflection] | 
| 54736 | 44 | else [Thm.symmetric (thm RS eq_reflection)]) | 
| 45 | handle Match => []; | |
| 54738 | 46 | val simpset = | 
| 47 | empty_simpset ctxt | |
| 54736 | 48 | |> Simplifier.set_mksimps (K mk_eq); | 
| 53666 | 49 | in | 
| 54738 | 50 | asm_lr_simp_tac simpset i | 
| 53666 | 51 | end; | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 52 | |
| 59652 | 53 | val eq_boolI = @{lemma "\<And>P. P \<Longrightarrow> P = True" "\<And>P. \<not> P \<Longrightarrow> P = False" by iprover+};
 | 
| 54736 | 54 | val boolE = @{thms HOL.TrueE HOL.FalseE};
 | 
| 59652 | 55 | val boolD = @{lemma "\<And>P. True = P \<Longrightarrow> P" "\<And>P. False = P \<Longrightarrow> \<not> P" by iprover+};
 | 
| 54737 | 56 | val eq_bool = @{thms HOL.eq_True HOL.eq_False HOL.not_False_eq_True HOL.not_True_eq_False};
 | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 57 | |
| 65136 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 58 | fun bool_subst_tac ctxt = | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 59 | TRY o REPEAT_ALL_NEW (EqSubst.eqsubst_asm_tac ctxt [1] eq_bool) | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 60 | THEN_ALL_NEW TRY o REPEAT_ALL_NEW (dresolve_tac ctxt boolD) | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 61 | THEN_ALL_NEW TRY o REPEAT_ALL_NEW (eresolve_tac ctxt boolE); | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 62 | |
| 53666 | 63 | fun mk_bool_elims ctxt elim = | 
| 54736 | 64 | let | 
| 65 | fun mk_bool_elim b = | |
| 66 | elim | |
| 67 | |> Thm.forall_elim b | |
| 65136 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 68 | |> Tactic.rule_by_tactic ctxt (distinct_subgoals_tac THEN TRY (resolve_tac ctxt eq_boolI 1)) | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 69 | |> Tactic.rule_by_tactic ctxt (distinct_subgoals_tac THEN ALLGOALS (bool_subst_tac ctxt)); | 
| 53666 | 70 | in | 
| 67149 | 71 | map mk_bool_elim [\<^cterm>\<open>True\<close>, \<^cterm>\<open>False\<close>] | 
| 53666 | 72 | end; | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 73 | |
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 74 | in | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 75 | |
| 53667 
0aefb31e27e0
distinguish Proof.context vs. local_theory semantically, with corresponding naming conventions;
 wenzelm parents: 
53666diff
changeset | 76 | fun mk_partial_elim_rules ctxt result = | 
| 
0aefb31e27e0
distinguish Proof.context vs. local_theory semantically, with corresponding naming conventions;
 wenzelm parents: 
53666diff
changeset | 77 | let | 
| 59655 | 78 |     val Function_Common.FunctionResult {fs, R, dom, psimps, cases, ...} = result;
 | 
| 54736 | 79 | val n_fs = length fs; | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 80 | |
| 59653 | 81 | fun variant_free used_term v = | 
| 82 | Free (singleton (Variable.variant_frees ctxt [used_term]) v); | |
| 83 | ||
| 54736 | 84 | fun mk_partial_elim_rule (idx, f) = | 
| 85 | let | |
| 59653 | 86 | fun mk_funeq 0 T (acc_args, acc_lhs) = | 
| 87 |               let val y = variant_free acc_lhs ("y", T)
 | |
| 88 | in (y, rev acc_args, HOLogic.mk_Trueprop (HOLogic.mk_eq (acc_lhs, y))) end | |
| 67149 | 89 | | mk_funeq n (Type (\<^type_name>\<open>fun\<close>, [S, T])) (acc_args, acc_lhs) = | 
| 59653 | 90 |               let val x = variant_free acc_lhs ("x", S)
 | 
| 91 | in mk_funeq (n - 1) T (x :: acc_args, acc_lhs $ x) end | |
| 92 |           | mk_funeq _ _ _ = raise TERM ("Not a function", [f]);
 | |
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 93 | |
| 54736 | 94 | val f_simps = | 
| 95 | filter (fn r => | |
| 59652 | 96 | (Thm.prop_of r | 
| 97 | |> Logic.strip_assums_concl | |
| 54736 | 98 | |> HOLogic.dest_Trueprop | 
| 99 | |> dest_funprop |> fst |> fst) = f) | |
| 100 | psimps; | |
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 101 | |
| 54736 | 102 | val arity = | 
| 103 | hd f_simps | |
| 59582 | 104 | |> Thm.prop_of | 
| 54736 | 105 | |> Logic.strip_assums_concl | 
| 106 | |> HOLogic.dest_Trueprop | |
| 107 | |> snd o fst o dest_funprop | |
| 108 | |> length; | |
| 59652 | 109 | |
| 59653 | 110 | val (rhs_var, arg_vars, prop) = mk_funeq arity (fastype_of f) ([], f); | 
| 54736 | 111 | val args = HOLogic.mk_tuple arg_vars; | 
| 112 | val domT = R |> dest_Free |> snd |> hd o snd o dest_Type; | |
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 113 | |
| 67149 | 114 |         val P = Thm.cterm_of ctxt (variant_free prop ("P", \<^typ>\<open>bool\<close>));
 | 
| 59653 | 115 | val sumtree_inj = Sum_Tree.mk_inj domT n_fs (idx + 1) args; | 
| 54736 | 116 | |
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 117 | val cprop = Thm.cterm_of ctxt prop; | 
| 54736 | 118 | |
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 119 | val asms = [cprop, Thm.cterm_of ctxt (HOLogic.mk_Trueprop (dom $ sumtree_inj))]; | 
| 54736 | 120 | val asms_thms = map Thm.assume asms; | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 121 | |
| 65136 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 122 | val prep_subgoal_tac = | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 123 |           TRY o REPEAT_ALL_NEW (eresolve_tac ctxt @{thms Pair_inject})
 | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 124 | THEN' Method.insert_tac ctxt | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 125 | (case asms_thms of thm :: thms => (thm RS sym) :: thms) | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 126 | THEN' propagate_tac ctxt | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 127 | THEN_ALL_NEW | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 128 | ((TRY o (EqSubst.eqsubst_asm_tac ctxt [1] psimps THEN' assume_tac ctxt))) | 
| 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 129 | THEN_ALL_NEW bool_subst_tac ctxt; | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 130 | |
| 59454 | 131 | val elim_stripped = | 
| 132 | nth cases idx | |
| 133 | |> Thm.forall_elim P | |
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 134 | |> Thm.forall_elim (Thm.cterm_of ctxt args) | 
| 65136 
115bcddf2ea2
Tuned generation of elimination rules in function package
 eberlm <eberlm@in.tum.de> parents: 
61841diff
changeset | 135 | |> Tactic.rule_by_tactic ctxt (distinct_subgoals_tac THEN ALLGOALS prep_subgoal_tac) | 
| 59454 | 136 | |> fold_rev Thm.implies_intr asms | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 137 | |> Thm.forall_intr (Thm.cterm_of ctxt rhs_var); | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 138 | |
| 59454 | 139 | val bool_elims = | 
| 67149 | 140 | if fastype_of rhs_var = \<^typ>\<open>bool\<close> | 
| 59655 | 141 | then mk_bool_elims ctxt elim_stripped | 
| 142 | else []; | |
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 143 | |
| 59454 | 144 | fun unstrip rl = | 
| 145 | rl | |
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 146 | |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) arg_vars | 
| 59454 | 147 | |> Thm.forall_intr P | 
| 148 | in | |
| 70479 | 149 | map (unstrip #> Thm.solve_constraints) (elim_stripped :: bool_elims) | 
| 59454 | 150 | end; | 
| 53666 | 151 | in | 
| 152 | map_index mk_partial_elim_rule fs | |
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 153 | end; | 
| 53666 | 154 | |
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 155 | end; | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 156 | |
| 53666 | 157 | end; |