src/ZF/Datatype.ML
author skalberg
Sun, 13 Feb 2005 17:15:14 +0100
changeset 15531 08c8dad8e399
parent 13780 af7b79271364
child 15570 8d8c70b41bab
permissions -rw-r--r--
Deleted Library.option type.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
     1
(*  Title:      ZF/Datatype.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 70
diff changeset
     6
(Co)Datatype Definitions for Zermelo-Fraenkel Set Theory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    10
(*Typechecking rules for most datatypes involving univ*)
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    11
structure Data_Arg =
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    12
  struct
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    13
  val intrs = 
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    14
      [SigmaI, InlI, InrI,
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    15
       Pair_in_univ, Inl_in_univ, Inr_in_univ, 
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    16
       zero_in_univ, A_into_univ, nat_into_univ, UnCI];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
6112
5e4871c5136b datatype package improvements
paulson
parents: 6053
diff changeset
    18
5e4871c5136b datatype package improvements
paulson
parents: 6053
diff changeset
    19
  val elims = [make_elim InlD, make_elim InrD,   (*for mutual recursion*)
12134
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    20
               SigmaE, sumE];                    (*allows * and + in spec*)
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    21
  end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    22
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    23
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    24
structure Data_Package = 
12183
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    25
  Add_datatype_def_Fun
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    26
   (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    27
    and Su=Standard_Sum
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    28
    and Ind_Package = Ind_Package
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    29
    and Datatype_Arg = Data_Arg
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    30
    val coind = false);
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    31
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    32
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    33
(*Typechecking rules for most codatatypes involving quniv*)
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    34
structure CoData_Arg =
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    35
  struct
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    36
  val intrs = 
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    37
      [QSigmaI, QInlI, QInrI,
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    38
       QPair_in_quniv, QInl_in_quniv, QInr_in_quniv, 
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    39
       zero_in_quniv, A_into_quniv, nat_into_quniv, UnCI];
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    40
6112
5e4871c5136b datatype package improvements
paulson
parents: 6053
diff changeset
    41
  val elims = [make_elim QInlD, make_elim QInrD,   (*for mutual recursion*)
12134
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    42
               QSigmaE, qsumE];                    (*allows * and + in spec*)
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    43
  end;
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    44
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 1461
diff changeset
    45
structure CoData_Package = 
12183
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    46
  Add_datatype_def_Fun
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    47
   (structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    48
    and Su=Quine_Sum
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    49
    and Ind_Package = CoInd_Package
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    50
    and Datatype_Arg = CoData_Arg
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    51
    val coind = true);
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    52
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    53
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    54
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    55
(*Simproc for freeness reasoning: compare datatype constructors for equality*)
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    56
structure DataFree =
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    57
struct
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    58
  val trace = ref false;
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    59
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    60
  fun mk_new ([],[]) = Const("True",FOLogic.oT)
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    61
    | mk_new (largs,rargs) =
12134
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    62
        fold_bal FOLogic.mk_conj
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    63
                 (map FOLogic.mk_eq (ListPair.zip (largs,rargs)));
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    64
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12134
diff changeset
    65
 val datatype_ss = simpset_of (the_context ());
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    66
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    67
 fun proc sg _ old =
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    68
   let val _ = if !trace then writeln ("data_free: OLD = " ^ 
12134
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    69
                                       string_of_cterm (cterm_of sg old))
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    70
               else ()
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    71
       val (lhs,rhs) = FOLogic.dest_eq old
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    72
       val (lhead, largs) = strip_comb lhs
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    73
       and (rhead, rargs) = strip_comb rhs
12203
571d9c288640 avoid handle _;
wenzelm
parents: 12183
diff changeset
    74
       val lname = #1 (dest_Const lhead) handle TERM _ => raise Match;
571d9c288640 avoid handle _;
wenzelm
parents: 12183
diff changeset
    75
       val rname = #1 (dest_Const rhead) handle TERM _ => raise Match;
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    76
       val lcon_info = the (Symtab.lookup (ConstructorsData.get_sg sg, lname))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13780
diff changeset
    77
         handle Option => raise Match;
12203
571d9c288640 avoid handle _;
wenzelm
parents: 12183
diff changeset
    78
       val rcon_info = the (Symtab.lookup (ConstructorsData.get_sg sg, rname))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13780
diff changeset
    79
         handle Option => raise Match;
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    80
       val new = 
12134
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    81
           if #big_rec_name lcon_info = #big_rec_name rcon_info 
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    82
               andalso not (null (#free_iffs lcon_info)) then
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    83
               if lname = rname then mk_new (largs, rargs)
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    84
               else Const("False",FOLogic.oT)
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    85
           else raise Match
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    86
       val _ = if !trace then 
12134
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    87
                 writeln ("NEW = " ^ string_of_cterm (Thm.cterm_of sg new))
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    88
               else ();
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    89
       val goal = Logic.mk_equals (old, new)
7049eead7a50 use Tactic.prove;
wenzelm
parents: 9000
diff changeset
    90
       val thm = Tactic.prove sg [] [] goal (fn _ => rtac iff_reflection 1 THEN
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12134
diff changeset
    91
           simp_tac (datatype_ss addsimps #free_iffs lcon_info) 1)
12203
571d9c288640 avoid handle _;
wenzelm
parents: 12183
diff changeset
    92
         handle ERROR_MESSAGE msg =>
571d9c288640 avoid handle _;
wenzelm
parents: 12183
diff changeset
    93
         (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Sign.string_of_term sg goal);
571d9c288640 avoid handle _;
wenzelm
parents: 12183
diff changeset
    94
          raise Match)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13780
diff changeset
    95
   in SOME thm end
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13780
diff changeset
    96
   handle Match => NONE;
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    97
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    98
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    99
 val conv = 
13780
af7b79271364 more new-style theories
paulson
parents: 13462
diff changeset
   100
     Simplifier.simproc (Theory.sign_of (theory "ZF")) "data_free" 
af7b79271364 more new-style theories
paulson
parents: 13462
diff changeset
   101
                        ["(x::i) = y"] proc;
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
   102
end;
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
   103
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
   104
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
   105
Addsimprocs [DataFree.conv];