| author | mengj | 
| Fri, 28 Apr 2006 05:59:32 +0200 | |
| changeset 19491 | cd6c71c57f53 | 
| parent 18678 | dd0c569fa43d | 
| child 20049 | f48c4a3a34bc | 
| permissions | -rw-r--r-- | 
| 516 | 1 | (* Title: ZF/Datatype.ML | 
| 0 | 2 | ID: $Id$ | 
| 516 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 120 | 6 | (Co)Datatype Definitions for Zermelo-Fraenkel Set Theory | 
| 0 | 7 | *) | 
| 8 | ||
| 9 | ||
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 10 | (*Typechecking rules for most datatypes involving univ*) | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 11 | structure Data_Arg = | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 12 | struct | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 13 | val intrs = | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 14 | [SigmaI, InlI, InrI, | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 15 | Pair_in_univ, Inl_in_univ, Inr_in_univ, | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 16 | zero_in_univ, A_into_univ, nat_into_univ, UnCI]; | 
| 0 | 17 | |
| 6112 | 18 | |
| 19 | val elims = [make_elim InlD, make_elim InrD, (*for mutual recursion*) | |
| 12134 | 20 | SigmaE, sumE]; (*allows * and + in spec*) | 
| 516 | 21 | end; | 
| 22 | ||
| 23 | ||
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 24 | structure Data_Package = | 
| 12183 | 25 | Add_datatype_def_Fun | 
| 26 | (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP | |
| 27 | and Su=Standard_Sum | |
| 28 | and Ind_Package = Ind_Package | |
| 29 | and Datatype_Arg = Data_Arg | |
| 30 | val coind = false); | |
| 516 | 31 | |
| 32 | ||
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 33 | (*Typechecking rules for most codatatypes involving quniv*) | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 34 | structure CoData_Arg = | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 35 | struct | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 36 | val intrs = | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 37 | [QSigmaI, QInlI, QInrI, | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 38 | QPair_in_quniv, QInl_in_quniv, QInr_in_quniv, | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 39 | zero_in_quniv, A_into_quniv, nat_into_quniv, UnCI]; | 
| 516 | 40 | |
| 6112 | 41 | val elims = [make_elim QInlD, make_elim QInrD, (*for mutual recursion*) | 
| 12134 | 42 | QSigmaE, qsumE]; (*allows * and + in spec*) | 
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 43 | end; | 
| 516 | 44 | |
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1461diff
changeset | 45 | structure CoData_Package = | 
| 12183 | 46 | Add_datatype_def_Fun | 
| 47 | (structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP | |
| 48 | and Su=Quine_Sum | |
| 49 | and Ind_Package = CoInd_Package | |
| 50 | and Datatype_Arg = CoData_Arg | |
| 51 | val coind = true); | |
| 516 | 52 | |
| 6141 | 53 | |
| 54 | ||
| 55 | (*Simproc for freeness reasoning: compare datatype constructors for equality*) | |
| 56 | structure DataFree = | |
| 57 | struct | |
| 58 | val trace = ref false; | |
| 59 | ||
| 60 |   fun mk_new ([],[]) = Const("True",FOLogic.oT)
 | |
| 61 | | mk_new (largs,rargs) = | |
| 12134 | 62 | fold_bal FOLogic.mk_conj | 
| 63 | (map FOLogic.mk_eq (ListPair.zip (largs,rargs))); | |
| 6141 | 64 | |
| 16973 | 65 | val datatype_ss = simpset (); | 
| 6141 | 66 | |
| 16973 | 67 | fun proc sg ss old = | 
| 6141 | 68 |    let val _ = if !trace then writeln ("data_free: OLD = " ^ 
 | 
| 12134 | 69 | string_of_cterm (cterm_of sg old)) | 
| 70 | else () | |
| 6141 | 71 | val (lhs,rhs) = FOLogic.dest_eq old | 
| 72 | val (lhead, largs) = strip_comb lhs | |
| 73 | and (rhead, rargs) = strip_comb rhs | |
| 12203 | 74 | val lname = #1 (dest_Const lhead) handle TERM _ => raise Match; | 
| 75 | val rname = #1 (dest_Const rhead) handle TERM _ => raise Match; | |
| 17412 | 76 | val lcon_info = the (Symtab.lookup (ConstructorsData.get sg) lname) | 
| 15531 | 77 | handle Option => raise Match; | 
| 17412 | 78 | val rcon_info = the (Symtab.lookup (ConstructorsData.get sg) rname) | 
| 15531 | 79 | handle Option => raise Match; | 
| 6141 | 80 | val new = | 
| 12134 | 81 | if #big_rec_name lcon_info = #big_rec_name rcon_info | 
| 82 | andalso not (null (#free_iffs lcon_info)) then | |
| 83 | if lname = rname then mk_new (largs, rargs) | |
| 84 |                else Const("False",FOLogic.oT)
 | |
| 85 | else raise Match | |
| 6141 | 86 | val _ = if !trace then | 
| 12134 | 87 |                  writeln ("NEW = " ^ string_of_cterm (Thm.cterm_of sg new))
 | 
| 88 | else (); | |
| 89 | val goal = Logic.mk_equals (old, new) | |
| 17956 | 90 | val thm = Goal.prove sg [] [] goal (fn _ => rtac iff_reflection 1 THEN | 
| 17877 
67d5ab1cb0d8
Simplifier.inherit_context instead of Simplifier.inherit_bounds;
 wenzelm parents: 
17412diff
changeset | 91 | simp_tac (Simplifier.inherit_context ss datatype_ss addsimps #free_iffs lcon_info) 1) | 
| 18678 | 92 | handle ERROR msg => | 
| 12203 | 93 | (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Sign.string_of_term sg goal); | 
| 94 | raise Match) | |
| 15531 | 95 | in SOME thm end | 
| 96 | handle Match => NONE; | |
| 6141 | 97 | |
| 98 | ||
| 16973 | 99 | val conv = Simplifier.simproc (theory "ZF") "data_free" ["(x::i) = y"] proc; | 
| 100 | ||
| 6141 | 101 | end; | 
| 102 | ||
| 103 | ||
| 104 | Addsimprocs [DataFree.conv]; |