src/HOLCF/One.thy
author webertj
Tue, 11 Sep 2007 14:26:49 +0200
changeset 24567 4970fb01aa01
parent 18111 2b56f74fd605
child 25131 2c8caac48ade
permissions -rw-r--r--
typo fixed, dead link removed
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2275
diff changeset
     1
(*  Title:      HOLCF/One.thy
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     2
    ID:         $Id$
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2275
diff changeset
     3
    Author:     Oscar Slotosch
16070
4a83dd540b88 removed LICENCE note -- everything is subject to Isabelle licence as
wenzelm
parents: 15577
diff changeset
     4
4a83dd540b88 removed LICENCE note -- everything is subject to Isabelle licence as
wenzelm
parents: 15577
diff changeset
     5
The unit domain.
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     6
*)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     7
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     8
header {* The unit domain *}
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     9
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    10
theory One
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    11
imports Lift
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    12
begin
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    13
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    14
types one = "unit lift"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    15
3717
e28553315355 eliminated rules;
wenzelm
parents: 2640
diff changeset
    16
constdefs
e28553315355 eliminated rules;
wenzelm
parents: 2640
diff changeset
    17
  ONE :: "one"
16747
934b6b36d794 cleaned up
huffman
parents: 16070
diff changeset
    18
  "ONE \<equiv> Def ()"
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2275
diff changeset
    19
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2275
diff changeset
    20
translations
3717
e28553315355 eliminated rules;
wenzelm
parents: 2640
diff changeset
    21
  "one" <= (type) "unit lift" 
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    22
16747
934b6b36d794 cleaned up
huffman
parents: 16070
diff changeset
    23
text {* Exhaustion and Elimination for type @{typ one} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    24
16747
934b6b36d794 cleaned up
huffman
parents: 16070
diff changeset
    25
lemma Exh_one: "t = \<bottom> \<or> t = ONE"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    26
apply (unfold ONE_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    27
apply (induct t)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    28
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    29
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    30
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    31
16747
934b6b36d794 cleaned up
huffman
parents: 16070
diff changeset
    32
lemma oneE: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = ONE \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    33
apply (rule Exh_one [THEN disjE])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    34
apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    35
apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    36
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    37
16747
934b6b36d794 cleaned up
huffman
parents: 16070
diff changeset
    38
lemma dist_less_one [simp]: "\<not> ONE \<sqsubseteq> \<bottom>"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    39
apply (unfold ONE_def)
16747
934b6b36d794 cleaned up
huffman
parents: 16070
diff changeset
    40
apply simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    41
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    42
16747
934b6b36d794 cleaned up
huffman
parents: 16070
diff changeset
    43
lemma dist_eq_one [simp]: "ONE \<noteq> \<bottom>" "\<bottom> \<noteq> ONE"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    44
apply (unfold ONE_def)
16747
934b6b36d794 cleaned up
huffman
parents: 16070
diff changeset
    45
apply simp_all
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    46
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    47
17838
3032e90c4975 added compactness theorems
huffman
parents: 16747
diff changeset
    48
lemma compact_ONE [simp]: "compact ONE"
3032e90c4975 added compactness theorems
huffman
parents: 16747
diff changeset
    49
by (rule compact_chfin)
3032e90c4975 added compactness theorems
huffman
parents: 16747
diff changeset
    50
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    51
text {* Case analysis function for type @{typ one} *}
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    52
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    53
constdefs
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    54
  one_when :: "'a::pcpo \<rightarrow> one \<rightarrow> 'a"
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    55
  "one_when \<equiv> \<Lambda> a. strictify\<cdot>(\<Lambda> _. a)"
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    56
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    57
translations
18111
2b56f74fd605 add case syntax for type one
huffman
parents: 18080
diff changeset
    58
  "case x of ONE \<Rightarrow> t" == "one_when\<cdot>t\<cdot>x"
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    59
  "\<Lambda> ONE. t" == "one_when\<cdot>t"
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    60
18111
2b56f74fd605 add case syntax for type one
huffman
parents: 18080
diff changeset
    61
lemma one_when1 [simp]: "(case \<bottom> of ONE \<Rightarrow> t) = \<bottom>"
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    62
by (simp add: one_when_def)
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    63
18111
2b56f74fd605 add case syntax for type one
huffman
parents: 18080
diff changeset
    64
lemma one_when2 [simp]: "(case ONE of ONE \<Rightarrow> t) = t"
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    65
by (simp add: one_when_def)
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    66
18111
2b56f74fd605 add case syntax for type one
huffman
parents: 18080
diff changeset
    67
lemma one_when3 [simp]: "(case x of ONE \<Rightarrow> ONE) = x"
18080
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    68
by (rule_tac p=x in oneE, simp_all)
c1a7490ee3ff add constant one_when; LAM pattern for ONE
huffman
parents: 17838
diff changeset
    69
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    70
end