src/HOLCF/Discrete.thy
author wenzelm
Thu, 01 Oct 2009 16:09:47 +0200
changeset 32813 dac196e23093
parent 31076 99fe356cbbc2
child 35900 aa5dfb03eb1e
permissions -rw-r--r--
handle Pattern.MATCH, not arbitrary exceptions; misc tuning, trying to reduce line length a bit;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
d364491ba718 add header
huffman
parents: 15555
diff changeset
     5
header {* Discrete cpo types *}
d364491ba718 add header
huffman
parents: 15555
diff changeset
     6
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
     7
theory Discrete
19105
3aabd46340e0 use minimal imports
huffman
parents: 16213
diff changeset
     8
imports Cont
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
     9
begin
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    10
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    11
datatype 'a discr = Discr "'a :: type"
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    12
26025
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25921
diff changeset
    13
subsection {* Type @{typ "'a discr"} is a discrete cpo *}
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
c00823ce7288 new-style class instantiation
huffman
parents: 25827
diff changeset
    16
begin
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    17
25902
c00823ce7288 new-style class instantiation
huffman
parents: 25827
diff changeset
    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
c00823ce7288 new-style class instantiation
huffman
parents: 25827
diff changeset
    20
    "(op \<sqsubseteq> :: 'a discr \<Rightarrow> 'a discr \<Rightarrow> bool) = (op =)"
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    21
26025
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25921
diff changeset
    22
instance ..
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25921
diff changeset
    23
end
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
    24
26025
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25921
diff changeset
    25
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
    26
by intro_classes (simp add: below_discr_def)
25902
c00823ce7288 new-style class instantiation
huffman
parents: 25827
diff changeset
    27
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
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
    29
by simp (* FIXME: same discrete_cpo - remove? is [iff] important? *)
25902
c00823ce7288 new-style class instantiation
huffman
parents: 25827
diff changeset
    30
15590
17f4f5afcd5f added subsection headings, cleaned up some proofs
huffman
parents: 15578
diff changeset
    31
subsection {* Type @{typ "'a discr"} is a cpo *}
17f4f5afcd5f added subsection headings, cleaned up some proofs
huffman
parents: 15578
diff changeset
    32
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    33
lemma discr_chain0: 
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    34
 "!!S::nat=>('a::type)discr. chain S ==> S i = S 0"
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    35
apply (unfold chain_def)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    36
apply (induct_tac "i")
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    37
apply (rule refl)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    38
apply (erule subst)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    39
apply (rule sym)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    40
apply fast
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    41
done
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    42
15639
99ed5113783b cleaned up some proofs
huffman
parents: 15590
diff changeset
    43
lemma discr_chain_range0 [simp]: 
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    44
 "!!S::nat=>('a::type)discr. chain(S) ==> range(S) = {S 0}"
15639
99ed5113783b cleaned up some proofs
huffman
parents: 15590
diff changeset
    45
by (fast elim: discr_chain0)
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    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
0ca392ab7f37 change class axiom chfin to rule_format
huffman
parents: 25906
diff changeset
    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
15590
17f4f5afcd5f added subsection headings, cleaned up some proofs
huffman
parents: 15578
diff changeset
    63
subsection {* @{term 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
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25921
diff changeset
    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
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    71
26025
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25921
diff changeset
    72
lemma Discr_undiscr [simp]: "Discr (undiscr y) = y"
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25921
diff changeset
    73
by (induct y) simp
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25921
diff changeset
    74
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    75
lemma discr_chain_f_range0:
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    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
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    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
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25921
diff changeset
    80
by (rule cont_discrete_cpo)
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    81
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
    82
end