src/HOLCF/Discrete.thy
author huffman
Wed, 02 Mar 2005 22:57:08 +0100
changeset 15563 9e125b675253
parent 15555 9d4dbd18ff2d
child 15578 d364491ba718
permissions -rw-r--r--
converted to new-style theory
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
    ID:         $Id$
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
     5
12030
wenzelm
parents: 2841
diff changeset
     6
Discrete CPOs.
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
     7
*)
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
     8
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
     9
theory Discrete
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    10
imports Cont Datatype
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    11
begin
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    12
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    13
datatype 'a discr = Discr "'a :: type"
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    14
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    15
instance discr :: (type) sq_ord ..
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    16
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    17
defs (overloaded)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    18
less_discr_def: "((op <<)::('a::type)discr=>'a discr=>bool)  ==  op ="
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    19
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    20
lemma discr_less_eq [iff]: "((x::('a::type)discr) << y) = (x = y)"
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    21
apply (unfold less_discr_def)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    22
apply (rule refl)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    23
done
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    24
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    25
instance discr :: (type) po
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    26
proof
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    27
  fix x y z :: "'a discr"
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    28
  show "x << x" by simp
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    29
  { assume "x << y" and "y << x" thus "x = y" by simp }
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    30
  { assume "x << y" and "y << z" thus "x << z" by simp }
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    31
qed
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
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
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    43
lemma discr_chain_range0: 
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    44
 "!!S::nat=>('a::type)discr. chain(S) ==> range(S) = {S 0}"
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    45
apply (fast elim: discr_chain0)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    46
done
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    47
declare discr_chain_range0 [simp]
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    48
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    49
lemma discr_cpo: 
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    50
 "!!S. chain S ==> ? x::('a::type)discr. range(S) <<| x"
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    51
apply (unfold is_lub_def is_ub_def)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    52
apply (simp (no_asm_simp))
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    53
done
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    54
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    55
instance discr :: (type)cpo
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    56
by (intro_classes, rule discr_cpo)
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
    57
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
    58
constdefs
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    59
   undiscr :: "('a::type)discr => 'a"
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
    60
  "undiscr x == (case x of Discr y => y)"
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
    61
15555
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    62
lemma undiscr_Discr [simp]: "undiscr(Discr x) = x"
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    63
apply (unfold undiscr_def)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    64
apply (simp (no_asm))
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    65
done
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    66
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    67
lemma discr_chain_f_range0:
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    68
 "!!S::nat=>('a::type)discr. chain(S) ==> range(%i. f(S i)) = {f(S 0)}"
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    69
apply (fast dest: discr_chain0 elim: arg_cong)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    70
done
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    71
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    72
lemma cont_discr [iff]: "cont(%x::('a::type)discr. f x)"
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    73
apply (unfold cont is_lub_def is_ub_def)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    74
apply (simp (no_asm) add: discr_chain_f_range0)
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    75
done
9d4dbd18ff2d converted to new-style theory
huffman
parents: 14981
diff changeset
    76
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents:
diff changeset
    77
end