| author | wenzelm | 
| Thu, 06 Dec 2007 00:21:30 +0100 | |
| changeset 25551 | 87d89b0f847a | 
| parent 24893 | b8ef7afe3a6b | 
| permissions | -rw-r--r-- | 
| 12175 | 1  | 
(* Title: ZF/Datatype.thy  | 
| 
2870
 
6d6fd10a9fdc
Now a non-trivial theory so that require_thy can find it
 
paulson 
parents: 
809 
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
6d6fd10a9fdc
Now a non-trivial theory so that require_thy can find it
 
paulson 
parents: 
809 
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
6d6fd10a9fdc
Now a non-trivial theory so that require_thy can find it
 
paulson 
parents: 
809 
diff
changeset
 | 
4  | 
Copyright 1997 University of Cambridge  | 
| 
 
6d6fd10a9fdc
Now a non-trivial theory so that require_thy can find it
 
paulson 
parents: 
809 
diff
changeset
 | 
5  | 
|
| 
 
6d6fd10a9fdc
Now a non-trivial theory so that require_thy can find it
 
paulson 
parents: 
809 
diff
changeset
 | 
6  | 
*)  | 
| 516 | 7  | 
|
| 13328 | 8  | 
header{*Datatype and CoDatatype Definitions*}
 | 
9  | 
||
| 22814 | 10  | 
theory Datatype  | 
11  | 
imports Inductive Univ QUniv  | 
|
12  | 
uses "Tools/datatype_package.ML"  | 
|
13  | 
begin  | 
|
14  | 
||
15  | 
ML_setup {*
 | 
|
16  | 
(*Typechecking rules for most datatypes involving univ*)  | 
|
17  | 
structure Data_Arg =  | 
|
18  | 
struct  | 
|
19  | 
val intrs =  | 
|
| 24893 | 20  | 
      [@{thm SigmaI}, @{thm InlI}, @{thm InrI},
 | 
21  | 
       @{thm Pair_in_univ}, @{thm Inl_in_univ}, @{thm Inr_in_univ}, 
 | 
|
22  | 
       @{thm zero_in_univ}, @{thm A_into_univ}, @{thm nat_into_univ}, @{thm UnCI}];
 | 
|
| 22814 | 23  | 
|
24  | 
||
| 24893 | 25  | 
  val elims = [make_elim @{thm InlD}, make_elim @{thm InrD},   (*for mutual recursion*)
 | 
26  | 
               @{thm SigmaE}, @{thm sumE}];                    (*allows * and + in spec*)
 | 
|
| 22814 | 27  | 
end;  | 
28  | 
||
29  | 
||
30  | 
structure Data_Package =  | 
|
31  | 
Add_datatype_def_Fun  | 
|
32  | 
(structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP  | 
|
33  | 
and Su=Standard_Sum  | 
|
34  | 
and Ind_Package = Ind_Package  | 
|
35  | 
and Datatype_Arg = Data_Arg  | 
|
36  | 
val coind = false);  | 
|
37  | 
||
38  | 
||
39  | 
(*Typechecking rules for most codatatypes involving quniv*)  | 
|
40  | 
structure CoData_Arg =  | 
|
41  | 
struct  | 
|
42  | 
val intrs =  | 
|
| 24893 | 43  | 
      [@{thm QSigmaI}, @{thm QInlI}, @{thm QInrI},
 | 
44  | 
       @{thm QPair_in_quniv}, @{thm QInl_in_quniv}, @{thm QInr_in_quniv}, 
 | 
|
45  | 
       @{thm zero_in_quniv}, @{thm A_into_quniv}, @{thm nat_into_quniv}, @{thm UnCI}];
 | 
|
| 22814 | 46  | 
|
| 24893 | 47  | 
  val elims = [make_elim @{thm QInlD}, make_elim @{thm QInrD},   (*for mutual recursion*)
 | 
48  | 
               @{thm QSigmaE}, @{thm qsumE}];                    (*allows * and + in spec*)
 | 
|
| 22814 | 49  | 
end;  | 
50  | 
||
51  | 
structure CoData_Package =  | 
|
52  | 
Add_datatype_def_Fun  | 
|
53  | 
(structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP  | 
|
54  | 
and Su=Quine_Sum  | 
|
55  | 
and Ind_Package = CoInd_Package  | 
|
56  | 
and Datatype_Arg = CoData_Arg  | 
|
57  | 
val coind = true);  | 
|
58  | 
||
59  | 
||
60  | 
||
61  | 
(*Simproc for freeness reasoning: compare datatype constructors for equality*)  | 
|
62  | 
structure DataFree =  | 
|
63  | 
struct  | 
|
64  | 
val trace = ref false;  | 
|
65  | 
||
66  | 
  fun mk_new ([],[]) = Const("True",FOLogic.oT)
 | 
|
67  | 
| mk_new (largs,rargs) =  | 
|
| 23419 | 68  | 
BalancedTree.make FOLogic.mk_conj  | 
| 22814 | 69  | 
(map FOLogic.mk_eq (ListPair.zip (largs,rargs)));  | 
70  | 
||
71  | 
 val datatype_ss = @{simpset};
 | 
|
72  | 
||
73  | 
fun proc sg ss old =  | 
|
74  | 
   let val _ = if !trace then writeln ("data_free: OLD = " ^ 
 | 
|
75  | 
string_of_cterm (cterm_of sg old))  | 
|
76  | 
else ()  | 
|
77  | 
val (lhs,rhs) = FOLogic.dest_eq old  | 
|
78  | 
val (lhead, largs) = strip_comb lhs  | 
|
79  | 
and (rhead, rargs) = strip_comb rhs  | 
|
80  | 
val lname = #1 (dest_Const lhead) handle TERM _ => raise Match;  | 
|
81  | 
val rname = #1 (dest_Const rhead) handle TERM _ => raise Match;  | 
|
82  | 
val lcon_info = the (Symtab.lookup (ConstructorsData.get sg) lname)  | 
|
83  | 
handle Option => raise Match;  | 
|
84  | 
val rcon_info = the (Symtab.lookup (ConstructorsData.get sg) rname)  | 
|
85  | 
handle Option => raise Match;  | 
|
86  | 
val new =  | 
|
87  | 
if #big_rec_name lcon_info = #big_rec_name rcon_info  | 
|
88  | 
andalso not (null (#free_iffs lcon_info)) then  | 
|
89  | 
if lname = rname then mk_new (largs, rargs)  | 
|
90  | 
               else Const("False",FOLogic.oT)
 | 
|
91  | 
else raise Match  | 
|
92  | 
val _ = if !trace then  | 
|
93  | 
                 writeln ("NEW = " ^ string_of_cterm (Thm.cterm_of sg new))
 | 
|
94  | 
else ();  | 
|
95  | 
val goal = Logic.mk_equals (old, new)  | 
|
96  | 
val thm = Goal.prove (Simplifier.the_context ss) [] [] goal  | 
|
97  | 
(fn _ => rtac iff_reflection 1 THEN  | 
|
98  | 
simp_tac (Simplifier.inherit_context ss datatype_ss addsimps #free_iffs lcon_info) 1)  | 
|
99  | 
handle ERROR msg =>  | 
|
100  | 
(warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Sign.string_of_term sg goal);  | 
|
101  | 
raise Match)  | 
|
102  | 
in SOME thm end  | 
|
103  | 
handle Match => NONE;  | 
|
104  | 
||
105  | 
||
106  | 
 val conv = Simplifier.simproc @{theory} "data_free" ["(x::i) = y"] proc;
 | 
|
107  | 
||
108  | 
end;  | 
|
109  | 
||
110  | 
||
111  | 
Addsimprocs [DataFree.conv];  | 
|
112  | 
*}  | 
|
| 12175 | 113  | 
|
114  | 
end  |