0
|
1 |
(* Title: ZF/ex/llist.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1993 University of Cambridge
|
|
5 |
|
|
6 |
Co-Datatype definition of Lazy Lists
|
|
7 |
|
|
8 |
Needs a "take-lemma" to prove llist_subset_quniv and to justify co-induction
|
|
9 |
for proving equality
|
|
10 |
*)
|
|
11 |
|
|
12 |
structure LList = Co_Datatype_Fun
|
|
13 |
(val thy = QUniv.thy;
|
|
14 |
val thy = QUniv.thy;
|
|
15 |
val rec_specs =
|
|
16 |
[("llist", "quniv(A)",
|
|
17 |
[(["LNil"], "i"),
|
|
18 |
(["LCons"], "[i,i]=>i")])];
|
|
19 |
val rec_styp = "i=>i";
|
|
20 |
val ext = None
|
|
21 |
val sintrs =
|
|
22 |
["LNil : llist(A)",
|
|
23 |
"[| a: A; l: llist(A) |] ==> LCons(a,l) : llist(A)"];
|
|
24 |
val monos = [];
|
|
25 |
val type_intrs = co_data_typechecks
|
|
26 |
val type_elims = []);
|
|
27 |
|
|
28 |
val [LNilI, LConsI] = LList.intrs;
|
|
29 |
|
|
30 |
(*An elimination rule, for type-checking*)
|
|
31 |
val LConsE = LList.mk_cases LList.con_defs "LCons(a,l) : llist(A)";
|
|
32 |
|
|
33 |
(*Proving freeness results*)
|
|
34 |
val LCons_iff = LList.mk_free "LCons(a,l)=LCons(a',l') <-> a=a' & l=l'";
|
|
35 |
val LNil_LCons_iff = LList.mk_free "~ LNil=LCons(a,l)";
|
|
36 |
|
|
37 |
(*** Lemmas to justify using "llist" in other recursive type definitions ***)
|
|
38 |
|
|
39 |
goalw LList.thy LList.defs "!!A B. A<=B ==> llist(A) <= llist(B)";
|
|
40 |
by (rtac gfp_mono 1);
|
|
41 |
by (REPEAT (rtac LList.bnd_mono 1));
|
|
42 |
by (REPEAT (ares_tac (quniv_mono::basic_monos) 1));
|
|
43 |
val llist_mono = result();
|
|
44 |
|
|
45 |
(** Closure of quniv(A) under llist -- why so complex? Its a gfp... **)
|
|
46 |
|
|
47 |
val in_quniv_rls =
|
|
48 |
[Transset_quniv, QPair_Int_quniv_in_quniv, Int_Vfrom_0_in_quniv,
|
|
49 |
zero_Int_in_quniv, one_Int_in_quniv,
|
|
50 |
QPair_Int_Vfrom_succ_in_quniv, QPair_Int_Vfrom_in_quniv];
|
|
51 |
|
|
52 |
val quniv_cs = ZF_cs addSIs in_quniv_rls
|
|
53 |
addIs (Int_quniv_in_quniv::co_data_typechecks);
|
|
54 |
|
|
55 |
(*Keep unfolding the lazy list until the induction hypothesis applies*)
|
|
56 |
goal LList.thy
|
|
57 |
"!!i. i : nat ==> \
|
|
58 |
\ ALL l: llist(quniv(A)). l Int Vfrom(quniv(A), i) : quniv(A)";
|
|
59 |
be complete_induct 1;
|
|
60 |
br ballI 1;
|
|
61 |
be LList.elim 1;
|
|
62 |
bws ([QInl_def,QInr_def]@LList.con_defs);
|
|
63 |
by (fast_tac quniv_cs 1);
|
|
64 |
by (etac natE 1 THEN REPEAT_FIRST hyp_subst_tac);
|
|
65 |
by (fast_tac quniv_cs 1);
|
|
66 |
by (fast_tac quniv_cs 1);
|
|
67 |
val llist_quniv_lemma = result();
|
|
68 |
|
|
69 |
goal LList.thy "llist(quniv(A)) <= quniv(A)";
|
|
70 |
br subsetI 1;
|
|
71 |
br quniv_Int_Vfrom 1;
|
|
72 |
be (LList.dom_subset RS subsetD) 1;
|
|
73 |
by (REPEAT (ares_tac [llist_quniv_lemma RS bspec] 1));
|
|
74 |
val llist_quniv = result();
|
|
75 |
|
|
76 |
val llist_subset_quniv = standard
|
|
77 |
(llist_mono RS (llist_quniv RSN (2,subset_trans)));
|
|
78 |
|
|
79 |
(*** Equality for llist(A) as a greatest fixed point ***)
|
|
80 |
|
|
81 |
structure LList_Eq = Co_Inductive_Fun
|
|
82 |
(val thy = LList.thy addconsts [(["lleq"],"i=>i")];
|
|
83 |
val rec_doms = [("lleq","llist(A) <*> llist(A)")];
|
|
84 |
val sintrs =
|
|
85 |
["<LNil; LNil> : lleq(A)",
|
|
86 |
"[| a:A; <l; l'>: lleq(A) |] ==> <LCons(a,l); LCons(a,l')> : lleq(A)"];
|
|
87 |
val monos = [];
|
|
88 |
val con_defs = [];
|
|
89 |
val type_intrs = LList.intrs@[QSigmaI];
|
|
90 |
val type_elims = [QSigmaE2]);
|
|
91 |
|
|
92 |
(** Alternatives for above:
|
|
93 |
val con_defs = LList.con_defs
|
|
94 |
val type_intrs = co_data_typechecks
|
|
95 |
val type_elims = [quniv_QPair_E]
|
|
96 |
**)
|
|
97 |
|
|
98 |
val lleq_cs = subset_cs
|
|
99 |
addSIs [succI1, Int_Vset_0_subset,
|
|
100 |
QPair_Int_Vset_succ_subset_trans,
|
|
101 |
QPair_Int_Vset_subset_trans];
|
|
102 |
|
|
103 |
(*Keep unfolding the lazy list until the induction hypothesis applies*)
|
|
104 |
goal LList_Eq.thy
|
|
105 |
"!!i. Ord(i) ==> ALL l l'. <l;l'> : lleq(A) --> l Int Vset(i) <= l'";
|
|
106 |
be trans_induct 1;
|
|
107 |
by (safe_tac subset_cs);
|
|
108 |
be LList_Eq.elim 1;
|
|
109 |
by (safe_tac (subset_cs addSEs [QPair_inject]));
|
|
110 |
bws LList.con_defs;
|
|
111 |
by (etac Ord_cases 1 THEN REPEAT_FIRST hyp_subst_tac);
|
|
112 |
(*0 case*)
|
|
113 |
by (fast_tac lleq_cs 1);
|
|
114 |
(*succ(j) case*)
|
|
115 |
bw QInr_def;
|
|
116 |
by (fast_tac lleq_cs 1);
|
|
117 |
(*Limit(i) case*)
|
|
118 |
be (Limit_Vfrom_eq RS ssubst) 1;
|
|
119 |
br (Int_UN_distrib RS ssubst) 1;
|
|
120 |
by (fast_tac lleq_cs 1);
|
|
121 |
val lleq_Int_Vset_subset_lemma = result();
|
|
122 |
|
|
123 |
val lleq_Int_Vset_subset = standard
|
|
124 |
(lleq_Int_Vset_subset_lemma RS spec RS spec RS mp);
|
|
125 |
|
|
126 |
|
|
127 |
(*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*)
|
|
128 |
val [prem] = goal LList_Eq.thy "<l;l'> : lleq(A) ==> <l';l> : lleq(A)";
|
|
129 |
br (prem RS qconverseI RS LList_Eq.co_induct) 1;
|
|
130 |
br (LList_Eq.dom_subset RS qconverse_type) 1;
|
|
131 |
by (safe_tac qconverse_cs);
|
|
132 |
be LList_Eq.elim 1;
|
|
133 |
by (ALLGOALS (fast_tac qconverse_cs));
|
|
134 |
val lleq_symmetric = result();
|
|
135 |
|
|
136 |
goal LList_Eq.thy "!!l l'. <l;l'> : lleq(A) ==> l=l'";
|
|
137 |
br equalityI 1;
|
|
138 |
by (REPEAT (ares_tac [lleq_Int_Vset_subset RS Int_Vset_subset] 1
|
|
139 |
ORELSE etac lleq_symmetric 1));
|
|
140 |
val lleq_implies_equal = result();
|
|
141 |
|
|
142 |
val [eqprem,lprem] = goal LList_Eq.thy
|
|
143 |
"[| l=l'; l: llist(A) |] ==> <l;l'> : lleq(A)";
|
|
144 |
by (res_inst_tac [("X", "{<l;l>. l: llist(A)}")] LList_Eq.co_induct 1);
|
|
145 |
br (lprem RS RepFunI RS (eqprem RS subst)) 1;
|
|
146 |
by (safe_tac qpair_cs);
|
|
147 |
be LList.elim 1;
|
|
148 |
by (ALLGOALS (fast_tac qpair_cs));
|
|
149 |
val equal_llist_implies_leq = result();
|
|
150 |
|
|
151 |
|