src/ZF/Tools/ind_cases.ML
author wenzelm
Fri, 20 Dec 2019 15:24:34 +0100
changeset 71329 2217c731d228
parent 69593 3dda49e08b9d
child 74342 5d411d85da8c
permissions -rw-r--r--
tuned documentation;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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: 35762
diff changeset
     9
  val declare: string -> (Proof.context -> conv) -> theory -> theory
28084
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
    10
  val inductive_cases: (Attrib.binding * string list) list -> theory -> theory
12174
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    11
end;
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    12
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    13
structure IndCases: IND_CASES =
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    14
struct
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    15
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
(* theory data *)
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    18
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32149
diff changeset
    19
structure IndCasesData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22101
diff changeset
    20
(
36864
wenzelm
parents: 36610
diff changeset
    21
  type T = (Proof.context -> conv) Symtab.table;
12174
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    22
  val empty = Symtab.empty;
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 15705
diff changeset
    23
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32149
diff changeset
    24
  fun merge data = Symtab.merge (K true) data;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22101
diff changeset
    25
);
12174
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    26
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    27
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17224
diff changeset
    28
fun declare name f = IndCasesData.map (Symtab.update (name, f));
12174
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    29
36541
de1862c4a308 more explicit treatment of context, although not fully localized;
wenzelm
parents: 35762
diff changeset
    30
fun smart_cases ctxt s =
12174
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    31
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 39557
diff changeset
    32
    val thy = Proof_Context.theory_of ctxt;
59626
a6e977d8b070 clarified context;
wenzelm
parents: 59621
diff changeset
    33
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18418
diff changeset
    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: 35762
diff 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
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18418
diff changeset
    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
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17224
diff changeset
    39
    (case Symtab.lookup (IndCasesData.get thy) c of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15032
diff changeset
    40
      NONE => error ("Unknown inductive cases rule for set " ^ quote c)
59626
a6e977d8b070 clarified context;
wenzelm
parents: 59621
diff changeset
    41
    | SOME f => f ctxt (Thm.cterm_of ctxt 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
fb073a34b537 'inductive_cases': support 'and' form;
wenzelm
parents: 12311
diff changeset
    47
fun inductive_cases args thy =
12174
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    48
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 39557
diff changeset
    49
    val ctxt = Proof_Context.init_global thy;
12609
fb073a34b537 'inductive_cases': support 'and' form;
wenzelm
parents: 12311
diff changeset
    50
    val facts = args |> map (fn ((name, srcs), props) =>
56031
wenzelm
parents: 55111
diff changeset
    51
      ((name, map (Attrib.attribute_cmd ctxt) srcs),
36541
de1862c4a308 more explicit treatment of context, although not fully localized;
wenzelm
parents: 35762
diff 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: 36960
diff 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
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24508
diff changeset
    55
val _ =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63120
diff changeset
    56
  Outer_Syntax.command \<^command_keyword>\<open>inductive_cases\<close>
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 42361
diff changeset
    57
    "create simplified instances of elimination rules (improper)"
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
    58
    (Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Scan.repeat1 Parse.prop)
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24508
diff changeset
    59
      >> (Toplevel.theory o inductive_cases));
12174
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    60
58828
6d076fdd933d modernized setup;
wenzelm
parents: 56031
diff changeset
    61
6d076fdd933d modernized setup;
wenzelm
parents: 56031
diff changeset
    62
(* ind_cases method *)
6d076fdd933d modernized setup;
wenzelm
parents: 56031
diff changeset
    63
6d076fdd933d modernized setup;
wenzelm
parents: 56031
diff changeset
    64
val _ =
6d076fdd933d modernized setup;
wenzelm
parents: 56031
diff changeset
    65
  Theory.setup
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63120
diff changeset
    66
    (Method.setup \<^binding>\<open>ind_cases\<close>
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 59936
diff changeset
    67
      (Scan.lift (Scan.repeat1 Args.embedded_inner_syntax) >>
58828
6d076fdd933d modernized setup;
wenzelm
parents: 56031
diff changeset
    68
        (fn props => fn ctxt => Method.erule ctxt 0 (map (smart_cases ctxt) props)))
6d076fdd933d modernized setup;
wenzelm
parents: 56031
diff changeset
    69
      "dynamic case analysis on sets");
6d076fdd933d modernized setup;
wenzelm
parents: 56031
diff changeset
    70
12174
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    71
end;
a0aab0b9f2e9 Generic inductive cases facility for (co)inductive definitions.
wenzelm
parents:
diff changeset
    72