| author | blanchet | 
| Fri, 23 Apr 2010 12:24:30 +0200 | |
| changeset 36294 | 59a55dfa76d5 | 
| parent 35900 | aa5dfb03eb1e | 
| child 40089 | 8adc57fb8454 | 
| permissions | -rw-r--r-- | 
| 
2841
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOLCF/Discrete.thy  | 
| 
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
2  | 
Author: Tobias Nipkow  | 
| 
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
4  | 
|
| 15578 | 5  | 
header {* Discrete cpo types *}
 | 
6  | 
||
| 15555 | 7  | 
theory Discrete  | 
| 19105 | 8  | 
imports Cont  | 
| 15555 | 9  | 
begin  | 
10  | 
||
11  | 
datatype 'a discr = Discr "'a :: type"  | 
|
12  | 
||
| 
35900
 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 
huffman 
parents: 
31076 
diff
changeset
 | 
13  | 
subsection {* Discrete ordering *}
 | 
| 
15590
 
17f4f5afcd5f
added subsection headings, cleaned up some proofs
 
huffman 
parents: 
15578 
diff
changeset
 | 
14  | 
|
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
29138 
diff
changeset
 | 
15  | 
instantiation discr :: (type) below  | 
| 25902 | 16  | 
begin  | 
| 15555 | 17  | 
|
| 25902 | 18  | 
definition  | 
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
29138 
diff
changeset
 | 
19  | 
below_discr_def:  | 
| 25902 | 20  | 
"(op \<sqsubseteq> :: 'a discr \<Rightarrow> 'a discr \<Rightarrow> bool) = (op =)"  | 
| 15555 | 21  | 
|
| 26025 | 22  | 
instance ..  | 
23  | 
end  | 
|
| 
2841
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
24  | 
|
| 
35900
 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 
huffman 
parents: 
31076 
diff
changeset
 | 
25  | 
subsection {* Discrete cpo class instance *}
 | 
| 
 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 
huffman 
parents: 
31076 
diff
changeset
 | 
26  | 
|
| 26025 | 27  | 
instance discr :: (type) discrete_cpo  | 
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
29138 
diff
changeset
 | 
28  | 
by intro_classes (simp add: below_discr_def)  | 
| 25902 | 29  | 
|
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
29138 
diff
changeset
 | 
30  | 
lemma discr_below_eq [iff]: "((x::('a::type)discr) << y) = (x = y)"
 | 
| 
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
29138 
diff
changeset
 | 
31  | 
by simp (* FIXME: same discrete_cpo - remove? is [iff] important? *)  | 
| 25902 | 32  | 
|
| 15555 | 33  | 
lemma discr_chain0:  | 
34  | 
 "!!S::nat=>('a::type)discr. chain S ==> S i = S 0"
 | 
|
35  | 
apply (unfold chain_def)  | 
|
36  | 
apply (induct_tac "i")  | 
|
37  | 
apply (rule refl)  | 
|
38  | 
apply (erule subst)  | 
|
39  | 
apply (rule sym)  | 
|
40  | 
apply fast  | 
|
41  | 
done  | 
|
42  | 
||
| 15639 | 43  | 
lemma discr_chain_range0 [simp]:  | 
| 15555 | 44  | 
 "!!S::nat=>('a::type)discr. chain(S) ==> range(S) = {S 0}"
 | 
| 15639 | 45  | 
by (fast elim: discr_chain0)  | 
| 15555 | 46  | 
|
| 
25827
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
47  | 
instance discr :: (finite) finite_po  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
48  | 
proof  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
49  | 
have "finite (Discr ` (UNIV :: 'a set))"  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
50  | 
by (rule finite_imageI [OF finite])  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
51  | 
also have "(Discr ` (UNIV :: 'a set)) = UNIV"  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
52  | 
by (auto, case_tac x, auto)  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
53  | 
finally show "finite (UNIV :: 'a discr set)" .  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
54  | 
qed  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
55  | 
|
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
56  | 
instance discr :: (type) chfin  | 
| 25921 | 57  | 
apply intro_classes  | 
| 
25827
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
58  | 
apply (rule_tac x=0 in exI)  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
59  | 
apply (unfold max_in_chain_def)  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
60  | 
apply (clarify, erule discr_chain0 [symmetric])  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
61  | 
done  | 
| 
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
62  | 
|
| 
35900
 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 
huffman 
parents: 
31076 
diff
changeset
 | 
63  | 
subsection {* \emph{undiscr} *}
 | 
| 
2841
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
64  | 
|
| 
25131
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19105 
diff
changeset
 | 
65  | 
definition  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19105 
diff
changeset
 | 
66  | 
  undiscr :: "('a::type)discr => 'a" where
 | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19105 
diff
changeset
 | 
67  | 
"undiscr x = (case x of Discr y => y)"  | 
| 
2841
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
68  | 
|
| 26025 | 69  | 
lemma undiscr_Discr [simp]: "undiscr (Discr x) = x"  | 
| 
15590
 
17f4f5afcd5f
added subsection headings, cleaned up some proofs
 
huffman 
parents: 
15578 
diff
changeset
 | 
70  | 
by (simp add: undiscr_def)  | 
| 15555 | 71  | 
|
| 26025 | 72  | 
lemma Discr_undiscr [simp]: "Discr (undiscr y) = y"  | 
73  | 
by (induct y) simp  | 
|
74  | 
||
| 15555 | 75  | 
lemma discr_chain_f_range0:  | 
76  | 
 "!!S::nat=>('a::type)discr. chain(S) ==> range(%i. f(S i)) = {f(S 0)}"
 | 
|
| 
15590
 
17f4f5afcd5f
added subsection headings, cleaned up some proofs
 
huffman 
parents: 
15578 
diff
changeset
 | 
77  | 
by (fast dest: discr_chain0 elim: arg_cong)  | 
| 15555 | 78  | 
|
| 
25827
 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 
huffman 
parents: 
25782 
diff
changeset
 | 
79  | 
lemma cont_discr [iff]: "cont (%x::('a::type)discr. f x)"
 | 
| 26025 | 80  | 
by (rule cont_discrete_cpo)  | 
| 15555 | 81  | 
|
| 
2841
 
c2508f4ab739
Added "discrete" CPOs and modified IMP to use those rather than "lift"
 
nipkow 
parents:  
diff
changeset
 | 
82  | 
end  |