35762
|
1 |
(* Title: ZF/ex/CoUnit.thy
|
1478
|
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
515
|
3 |
Copyright 1994 University of Cambridge
|
|
4 |
*)
|
|
5 |
|
12228
|
6 |
header {* Trivial codatatype definitions, one of which goes wrong! *}
|
|
7 |
|
16417
|
8 |
theory CoUnit imports Main begin
|
515
|
9 |
|
12228
|
10 |
text {*
|
|
11 |
See discussion in: L C Paulson. A Concrete Final Coalgebra Theorem
|
|
12 |
for ZF Set Theory. Report 334, Cambridge University Computer
|
|
13 |
Laboratory. 1994.
|
|
14 |
|
|
15 |
\bigskip
|
|
16 |
|
|
17 |
This degenerate definition does not work well because the one
|
|
18 |
constructor's definition is trivial! The same thing occurs with
|
|
19 |
Aczel's Special Final Coalgebra Theorem.
|
|
20 |
*}
|
|
21 |
|
515
|
22 |
consts
|
1401
|
23 |
counit :: i
|
515
|
24 |
codatatype
|
12228
|
25 |
"counit" = Con ("x \<in> counit")
|
|
26 |
|
|
27 |
inductive_cases ConE: "Con(x) \<in> counit"
|
|
28 |
-- {* USELESS because folding on @{term "Con(xa) == xa"} fails. *}
|
|
29 |
|
|
30 |
lemma Con_iff: "Con(x) = Con(y) <-> x = y"
|
|
31 |
-- {* Proving freeness results. *}
|
|
32 |
by (auto elim!: counit.free_elims)
|
|
33 |
|
|
34 |
lemma counit_eq_univ: "counit = quniv(0)"
|
|
35 |
-- {* Should be a singleton, not everything! *}
|
|
36 |
apply (rule counit.dom_subset [THEN equalityI])
|
|
37 |
apply (rule subsetI)
|
|
38 |
apply (erule counit.coinduct)
|
|
39 |
apply (rule subset_refl)
|
|
40 |
apply (unfold counit.con_defs)
|
|
41 |
apply fast
|
|
42 |
done
|
515
|
43 |
|
|
44 |
|
12228
|
45 |
text {*
|
|
46 |
\medskip A similar example, but the constructor is non-degenerate
|
|
47 |
and it works! The resulting set is a singleton.
|
|
48 |
*}
|
515
|
49 |
|
|
50 |
consts
|
1401
|
51 |
counit2 :: i
|
515
|
52 |
codatatype
|
12228
|
53 |
"counit2" = Con2 ("x \<in> counit2", "y \<in> counit2")
|
|
54 |
|
|
55 |
|
|
56 |
inductive_cases Con2E: "Con2(x, y) \<in> counit2"
|
|
57 |
|
|
58 |
lemma Con2_iff: "Con2(x, y) = Con2(x', y') <-> x = x' & y = y'"
|
|
59 |
-- {* Proving freeness results. *}
|
|
60 |
by (fast elim!: counit2.free_elims)
|
|
61 |
|
|
62 |
lemma Con2_bnd_mono: "bnd_mono(univ(0), %x. Con2(x, x))"
|
|
63 |
apply (unfold counit2.con_defs)
|
|
64 |
apply (rule bnd_monoI)
|
|
65 |
apply (assumption | rule subset_refl QPair_subset_univ QPair_mono)+
|
|
66 |
done
|
|
67 |
|
|
68 |
lemma lfp_Con2_in_counit2: "lfp(univ(0), %x. Con2(x,x)) \<in> counit2"
|
|
69 |
apply (rule singletonI [THEN counit2.coinduct])
|
|
70 |
apply (rule qunivI [THEN singleton_subsetI])
|
|
71 |
apply (rule subset_trans [OF lfp_subset empty_subsetI [THEN univ_mono]])
|
|
72 |
apply (fast intro!: Con2_bnd_mono [THEN lfp_unfold])
|
|
73 |
done
|
|
74 |
|
|
75 |
lemma counit2_Int_Vset_subset [rule_format]:
|
|
76 |
"Ord(i) ==> \<forall>x y. x \<in> counit2 --> y \<in> counit2 --> x Int Vset(i) \<subseteq> y"
|
|
77 |
-- {* Lemma for proving finality. *}
|
|
78 |
apply (erule trans_induct)
|
42793
|
79 |
apply (tactic "safe_tac (put_claset subset_cs @{context})")
|
12228
|
80 |
apply (erule counit2.cases)
|
|
81 |
apply (erule counit2.cases)
|
|
82 |
apply (unfold counit2.con_defs)
|
42793
|
83 |
apply (tactic {* fast_tac (put_claset subset_cs @{context}
|
24893
|
84 |
addSIs [@{thm QPair_Int_Vset_subset_UN} RS @{thm subset_trans}, @{thm QPair_mono}]
|
|
85 |
addSEs [@{thm Ord_in_Ord}, @{thm Pair_inject}]) 1 *})
|
12228
|
86 |
done
|
|
87 |
|
|
88 |
lemma counit2_implies_equal: "[| x \<in> counit2; y \<in> counit2 |] ==> x = y"
|
|
89 |
apply (rule equalityI)
|
|
90 |
apply (assumption | rule conjI counit2_Int_Vset_subset [THEN Int_Vset_subset])+
|
|
91 |
done
|
|
92 |
|
|
93 |
lemma counit2_eq_univ: "counit2 = {lfp(univ(0), %x. Con2(x,x))}"
|
|
94 |
apply (rule equalityI)
|
|
95 |
apply (rule_tac [2] lfp_Con2_in_counit2 [THEN singleton_subsetI])
|
|
96 |
apply (rule subsetI)
|
|
97 |
apply (drule lfp_Con2_in_counit2 [THEN counit2_implies_equal])
|
|
98 |
apply (erule subst)
|
|
99 |
apply (rule singletonI)
|
|
100 |
done
|
515
|
101 |
|
|
102 |
end
|