src/ZF/Datatype.thy
author wenzelm
Thu, 19 Oct 2023 17:06:39 +0200
changeset 78800 0b3700d31758
parent 78790 8f4424187193
permissions -rw-r--r--
clarified signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68490
eb53f944c8cd simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
wenzelm
parents: 62913
diff changeset
     1
(*  Title:      ZF/Datatype.thy
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     3
    Copyright   1997  University of Cambridge
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     4
*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     5
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59498
diff changeset
     6
section\<open>Datatype and CoDatatype Definitions\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     7
68490
eb53f944c8cd simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
wenzelm
parents: 62913
diff changeset
     8
theory Datatype
eb53f944c8cd simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
wenzelm
parents: 62913
diff changeset
     9
imports Inductive Univ QUniv
46950
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 41777
diff changeset
    10
keywords "datatype" "codatatype" :: thy_decl
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    11
begin
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    12
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    13
ML_file \<open>Tools/datatype_package.ML\<close>
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46950
diff changeset
    14
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59498
diff changeset
    15
ML \<open>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    16
(*Typechecking rules for most datatypes involving univ*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    17
structure Data_Arg =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    18
  struct
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32010
diff changeset
    19
  val intrs =
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    20
      [@{thm SigmaI}, @{thm InlI}, @{thm InrI},
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32010
diff changeset
    21
       @{thm Pair_in_univ}, @{thm Inl_in_univ}, @{thm Inr_in_univ},
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    22
       @{thm zero_in_univ}, @{thm A_into_univ}, @{thm nat_into_univ}, @{thm UnCI}];
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    23
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    24
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    25
  val elims = [make_elim @{thm InlD}, make_elim @{thm InrD},   (*for mutual recursion*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    26
               @{thm SigmaE}, @{thm sumE}];                    (*allows * and + in spec*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    27
  end;
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    28
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    29
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32010
diff changeset
    30
structure Data_Package =
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    31
  Add_datatype_def_Fun
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    32
   (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    33
    and Su=Standard_Sum
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    34
    and Ind_Package = Ind_Package
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    35
    and Datatype_Arg = Data_Arg
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    36
    val coind = false);
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    37
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    38
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    39
(*Typechecking rules for most codatatypes involving quniv*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    40
structure CoData_Arg =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    41
  struct
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32010
diff changeset
    42
  val intrs =
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    43
      [@{thm QSigmaI}, @{thm QInlI}, @{thm QInrI},
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32010
diff changeset
    44
       @{thm QPair_in_quniv}, @{thm QInl_in_quniv}, @{thm QInr_in_quniv},
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    45
       @{thm zero_in_quniv}, @{thm A_into_quniv}, @{thm nat_into_quniv}, @{thm UnCI}];
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    46
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    47
  val elims = [make_elim @{thm QInlD}, make_elim @{thm QInrD},   (*for mutual recursion*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    48
               @{thm QSigmaE}, @{thm qsumE}];                    (*allows * and + in spec*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    49
  end;
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    50
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32010
diff changeset
    51
structure CoData_Package =
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    52
  Add_datatype_def_Fun
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    53
   (structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    54
    and Su=Quine_Sum
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    55
    and Ind_Package = CoInd_Package
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    56
    and Datatype_Arg = CoData_Arg
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    57
    val coind = true);
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    58
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    59
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    60
78790
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    61
(* Simproc for freeness reasoning: compare datatype constructors for equality *)
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    62
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    63
structure Data_Free:
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    64
sig
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    65
  val trace: bool Config.T
78800
0b3700d31758 clarified signature;
wenzelm
parents: 78790
diff changeset
    66
  val proc: Simplifier.proc
78790
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    67
end =
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    68
struct
78790
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    69
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    70
val trace = Attrib.setup_config_bool \<^binding>\<open>data_free_trace\<close> (K false);
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    71
78790
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    72
fun mk_new ([],[]) = \<^Const>\<open>True\<close>
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    73
  | mk_new (largs,rargs) =
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    74
      Balanced_Tree.make FOLogic.mk_conj
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    75
                 (map FOLogic.mk_eq (ListPair.zip (largs,rargs)));
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    76
78790
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    77
val datatype_ss = simpset_of \<^context>;
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    78
78790
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    79
fun proc ctxt ct =
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    80
  let
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    81
    val old = Thm.term_of ct
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    82
    val thy = Proof_Context.theory_of ctxt
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    83
    val _ =
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    84
      if Config.get ctxt trace then tracing ("data_free: OLD = " ^ Syntax.string_of_term ctxt old)
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    85
      else ()
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    86
    val (lhs,rhs) = FOLogic.dest_eq old
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    87
    val (lhead, largs) = strip_comb lhs
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    88
    and (rhead, rargs) = strip_comb rhs
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    89
    val lname = #1 (dest_Const lhead) handle TERM _ => raise Match;
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    90
    val rname = #1 (dest_Const rhead) handle TERM _ => raise Match;
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    91
    val lcon_info = the (Symtab.lookup (ConstructorsData.get thy) lname)
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    92
      handle Option.Option => raise Match;
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    93
    val rcon_info = the (Symtab.lookup (ConstructorsData.get thy) rname)
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    94
      handle Option.Option => raise Match;
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    95
    val new =
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    96
      if #big_rec_name lcon_info = #big_rec_name rcon_info
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    97
          andalso not (null (#free_iffs lcon_info)) then
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    98
        if lname = rname then mk_new (largs, rargs)
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
    99
        else \<^Const>\<open>False\<close>
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   100
      else raise Match;
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   101
     val _ =
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   102
       if Config.get ctxt trace then tracing ("NEW = " ^ Syntax.string_of_term ctxt new)
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   103
       else ();
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   104
     val goal = Logic.mk_equals (old, new);
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   105
     val thm = Goal.prove ctxt [] [] goal
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   106
       (fn _ => resolve_tac ctxt @{thms iff_reflection} 1 THEN
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   107
         simp_tac (put_simpset datatype_ss ctxt addsimps
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   108
          (map (Thm.transfer thy) (#free_iffs lcon_info))) 1)
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   109
       handle ERROR msg =>
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   110
       (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Syntax.string_of_term ctxt goal);
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   111
        raise Match)
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   112
  in SOME thm end
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   113
  handle Match => NONE;
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   114
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   115
end;
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59498
diff changeset
   116
\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   117
78790
8f4424187193 more standard ML setup;
wenzelm
parents: 74294
diff changeset
   118
simproc_setup data_free ("(x::i) = y") = \<open>fn _ => Data_Free.proc\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   119
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   120
end