| author | wenzelm | 
| Wed, 25 Jan 2023 20:42:24 +0100 | |
| changeset 77097 | 023273cf2651 | 
| parent 70318 | 9eff9e2dc177 | 
| permissions | -rw-r--r-- | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 1 | (* Title: HOL/Tools/Function/fun_cases.ML | 
| 53609 | 2 | Author: Manuel Eberl, TU Muenchen | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 3 | |
| 53993 | 4 | The fun_cases command for generating specialised elimination rules for | 
| 5 | function package functions. | |
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 6 | *) | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 7 | |
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 8 | signature FUN_CASES = | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 9 | sig | 
| 53993 | 10 | val mk_fun_cases: Proof.context -> term -> thm | 
| 11 | val fun_cases: (Attrib.binding * term list) list -> local_theory -> | |
| 53995 | 12 | (string * thm list) list * local_theory | 
| 53993 | 13 | val fun_cases_cmd: (Attrib.binding * string list) list -> local_theory -> | 
| 53995 | 14 | (string * thm list) list * local_theory | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 15 | end; | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 16 | |
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 17 | structure Fun_Cases : FUN_CASES = | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 18 | struct | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 19 | |
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 20 | fun mk_fun_cases ctxt prop = | 
| 53992 | 21 | let | 
| 22 | fun err () = | |
| 23 | error (Pretty.string_of (Pretty.block | |
| 24 | [Pretty.str "Proposition is not a function equation:", | |
| 25 | Pretty.fbrk, Syntax.pretty_term ctxt prop])); | |
| 26 | val ((f, _), _) = | |
| 27 | Function_Elims.dest_funprop (HOLogic.dest_Trueprop prop) | |
| 28 | handle TERM _ => err (); | |
| 29 | val info = Function.get_info ctxt f handle List.Empty => err (); | |
| 30 |     val {elims, pelims, is_partial, ...} = info;
 | |
| 31 | val elims = if is_partial then pelims else the elims; | |
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 32 | val cprop = Thm.cterm_of ctxt prop; | 
| 53992 | 33 | fun mk_elim rl = | 
| 53994 | 34 | Thm.implies_intr cprop | 
| 35 | (Tactic.rule_by_tactic ctxt (Inductive.mk_cases_tac ctxt) (Thm.assume cprop RS rl)) | |
| 70318 
9eff9e2dc177
proper Proof_Context.export_morphism corresponding to Proof_Context.augment (see 7f568724d67e);
 wenzelm parents: 
70308diff
changeset | 36 | |> singleton (Proof_Context.export (Proof_Context.augment prop ctxt) ctxt); | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 37 | in | 
| 53992 | 38 | (case get_first (try mk_elim) (flat elims) of | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 39 | SOME r => r | 
| 53992 | 40 | | NONE => err ()) | 
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 41 | end; | 
| 53666 | 42 | |
| 53991 
d8f7f3d998de
provide regular ML interface and use plain Syntax.read_prop/Syntax.check_prop (update by Manuel Eberl);
 wenzelm parents: 
53670diff
changeset | 43 | fun gen_fun_cases prep_att prep_prop args lthy = | 
| 53667 
0aefb31e27e0
distinguish Proof.context vs. local_theory semantically, with corresponding naming conventions;
 wenzelm parents: 
53666diff
changeset | 44 | let | 
| 53991 
d8f7f3d998de
provide regular ML interface and use plain Syntax.read_prop/Syntax.check_prop (update by Manuel Eberl);
 wenzelm parents: 
53670diff
changeset | 45 | val thmss = | 
| 
d8f7f3d998de
provide regular ML interface and use plain Syntax.read_prop/Syntax.check_prop (update by Manuel Eberl);
 wenzelm parents: 
53670diff
changeset | 46 | map snd args | 
| 58993 
302104d8366b
prefer independent parallel map where user input is processed -- avoid non-deterministic feedback in error situations;
 wenzelm parents: 
55997diff
changeset | 47 | |> burrow (grouped 10 Par_List.map_independent (mk_fun_cases lthy o prep_prop lthy)); | 
| 53991 
d8f7f3d998de
provide regular ML interface and use plain Syntax.read_prop/Syntax.check_prop (update by Manuel Eberl);
 wenzelm parents: 
53670diff
changeset | 48 | val facts = | 
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
53995diff
changeset | 49 | map2 (fn ((a, atts), _) => fn thms => ((a, map (prep_att lthy) atts), [(thms, [])])) | 
| 53991 
d8f7f3d998de
provide regular ML interface and use plain Syntax.read_prop/Syntax.check_prop (update by Manuel Eberl);
 wenzelm parents: 
53670diff
changeset | 50 | args thmss; | 
| 53995 | 51 | in lthy |> Local_Theory.notes facts end; | 
| 53670 | 52 | |
| 53993 | 53 | val fun_cases = gen_fun_cases (K I) Syntax.check_prop; | 
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
53995diff
changeset | 54 | val fun_cases_cmd = gen_fun_cases Attrib.check_src Syntax.read_prop; | 
| 53667 
0aefb31e27e0
distinguish Proof.context vs. local_theory semantically, with corresponding naming conventions;
 wenzelm parents: 
53666diff
changeset | 55 | |
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 56 | val _ = | 
| 67149 | 57 | Outer_Syntax.local_theory \<^command_keyword>\<open>fun_cases\<close> | 
| 53995 | 58 | "create simplified instances of elimination rules for function equations" | 
| 63064 
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
 wenzelm parents: 
59936diff
changeset | 59 | (Parse.and_list1 Parse_Spec.simple_specs >> (snd oo fun_cases_cmd)); | 
| 53666 | 60 | |
| 53603 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 61 | end; | 
| 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 Manuel Eberl parents: diff
changeset | 62 |