| author | wenzelm | 
| Thu, 20 Sep 2012 15:00:14 +0200 | |
| changeset 49468 | 8b06b99fb85c | 
| parent 47815 | 43f677b3ae91 | 
| child 54742 | 7a86358a3c0b | 
| permissions | -rw-r--r-- | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 1 | (* Title: ZF/Tools/ind_cases.ML | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 2 | Author: Markus Wenzel, LMU Muenchen | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 3 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 4 | Generic inductive cases facility for (co)inductive definitions. | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 5 | *) | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 6 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 7 | signature IND_CASES = | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 8 | sig | 
| 36541 
de1862c4a308
more explicit treatment of context, although not fully localized;
 wenzelm parents: 
35762diff
changeset | 9 | val declare: string -> (Proof.context -> conv) -> theory -> theory | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 10 | val inductive_cases: (Attrib.binding * string list) list -> theory -> theory | 
| 18708 | 11 | val setup: theory -> theory | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 12 | end; | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 13 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 14 | structure IndCases: IND_CASES = | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 15 | struct | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 16 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 17 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 18 | (* theory data *) | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 19 | |
| 33522 | 20 | structure IndCasesData = Theory_Data | 
| 22846 | 21 | ( | 
| 36864 | 22 | type T = (Proof.context -> conv) Symtab.table; | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 23 | val empty = Symtab.empty; | 
| 16424 | 24 | val extend = I; | 
| 33522 | 25 | fun merge data = Symtab.merge (K true) data; | 
| 22846 | 26 | ); | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 27 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 28 | |
| 17412 | 29 | fun declare name f = IndCasesData.map (Symtab.update (name, f)); | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 30 | |
| 36541 
de1862c4a308
more explicit treatment of context, although not fully localized;
 wenzelm parents: 
35762diff
changeset | 31 | fun smart_cases ctxt s = | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 32 | let | 
| 42361 | 33 | val thy = Proof_Context.theory_of ctxt; | 
| 18678 | 34 |     fun err msg = cat_error msg ("Malformed set membership statement: " ^ s);
 | 
| 36541 
de1862c4a308
more explicit treatment of context, although not fully localized;
 wenzelm parents: 
35762diff
changeset | 35 | val A = Syntax.read_prop ctxt s handle ERROR msg => err msg; | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 36 | val c = #1 (Term.dest_Const (Term.head_of (#2 (Ind_Syntax.dest_mem (FOLogic.dest_Trueprop | 
| 18678 | 37 | (Logic.strip_imp_concl A)))))) handle TERM _ => err ""; | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 38 | in | 
| 17412 | 39 | (case Symtab.lookup (IndCasesData.get thy) c of | 
| 15531 | 40 |       NONE => error ("Unknown inductive cases rule for set " ^ quote c)
 | 
| 36541 
de1862c4a308
more explicit treatment of context, although not fully localized;
 wenzelm parents: 
35762diff
changeset | 41 | | SOME f => f ctxt (Thm.cterm_of thy A)) | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 42 | end; | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 43 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 44 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 45 | (* inductive_cases command *) | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 46 | |
| 12609 | 47 | fun inductive_cases args thy = | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 48 | let | 
| 42361 | 49 | val ctxt = Proof_Context.init_global thy; | 
| 12609 | 50 | val facts = args |> map (fn ((name, srcs), props) => | 
| 47815 | 51 | ((name, map (Attrib.attribute_cmd_global thy) srcs), | 
| 36541 
de1862c4a308
more explicit treatment of context, although not fully localized;
 wenzelm parents: 
35762diff
changeset | 52 | map (Thm.no_attributes o single o smart_cases ctxt) props)); | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
36960diff
changeset | 53 | in thy |> Global_Theory.note_thmss "" facts |> snd end; | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 54 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 55 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 56 | (* ind_cases method *) | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 57 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 58 | val setup = | 
| 30722 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30515diff
changeset | 59 |   Method.setup @{binding "ind_cases"}
 | 
| 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30515diff
changeset | 60 | (Scan.lift (Scan.repeat1 Args.name_source) >> | 
| 36541 
de1862c4a308
more explicit treatment of context, although not fully localized;
 wenzelm parents: 
35762diff
changeset | 61 | (fn props => fn ctxt => Method.erule 0 (map (smart_cases ctxt) props))) | 
| 30722 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30515diff
changeset | 62 | "dynamic case analysis on sets"; | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 63 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 64 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 65 | (* outer syntax *) | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 66 | |
| 24867 | 67 | val _ = | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
42361diff
changeset | 68 |   Outer_Syntax.command @{command_spec "inductive_cases"}
 | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
42361diff
changeset | 69 | "create simplified instances of elimination rules (improper)" | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 70 | (Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Scan.repeat1 Parse.prop) | 
| 24867 | 71 | >> (Toplevel.theory o inductive_cases)); | 
| 12174 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 72 | |
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 73 | end; | 
| 
a0aab0b9f2e9
Generic inductive cases facility for (co)inductive definitions.
 wenzelm parents: diff
changeset | 74 |