src/HOL/HOLCF/Pcpo.thy
author wenzelm
Mon, 06 Sep 2021 12:23:06 +0200
changeset 74245 282cd3aa6cc6
parent 69597 ff784d5a5bfb
child 74305 28a582aa25dd
permissions -rw-r--r--
clarified modules;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41431
diff changeset
     1
(*  Title:      HOL/HOLCF/Pcpo.thy
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
     2
    Author:     Franz Regensburger
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2394
diff changeset
     3
*)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 15563
diff changeset
     4
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
     5
section \<open>Classes cpo and pcpo\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 15563
diff changeset
     6
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     7
theory Pcpo
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
     8
  imports Porder
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     9
begin
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    10
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    11
subsection \<open>Complete partial orders\<close>
15588
14e3228f18cc arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    12
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    13
text \<open>The class cpo of chain complete partial orders\<close>
15588
14e3228f18cc arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    14
29614
1f7b1b0df292 simplified handling of base sort, dropped axclass
haftmann
parents: 29138
diff changeset
    15
class cpo = po +
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
    16
  assumes cpo: "chain S \<Longrightarrow> \<exists>x. range S <<| x"
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
    17
begin
2394
91d8abf108be adaptions for symbol font
oheimb
parents: 2291
diff changeset
    18
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    19
text \<open>in cpo's everthing equal to THE lub has lub properties for every chain\<close>
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    20
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
    21
lemma cpo_lubI: "chain S \<Longrightarrow> range S <<| (\<Squnion>i. S i)"
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40500
diff changeset
    22
  by (fast dest: cpo elim: is_lub_lub)
26026
f9647c040b58 add lemma cpo_lubI
huffman
parents: 26023
diff changeset
    23
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
    24
lemma thelubE: "\<lbrakk>chain S; (\<Squnion>i. S i) = l\<rbrakk> \<Longrightarrow> range S <<| l"
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40500
diff changeset
    25
  by (blast dest: cpo intro: is_lub_lub)
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    26
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    27
text \<open>Properties of the lub\<close>
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    28
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
    29
lemma is_ub_thelub: "chain S \<Longrightarrow> S x \<sqsubseteq> (\<Squnion>i. S i)"
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40500
diff changeset
    30
  by (blast dest: cpo intro: is_lub_lub [THEN is_lub_rangeD1])
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    31
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    32
lemma is_lub_thelub: "\<lbrakk>chain S; range S <| x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x"
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40500
diff changeset
    33
  by (blast dest: cpo intro: is_lub_lub [THEN is_lubD2])
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    34
39969
0b8e19f588a4 new lemmas about lub
huffman
parents: 39302
diff changeset
    35
lemma lub_below_iff: "chain S \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x \<longleftrightarrow> (\<forall>i. S i \<sqsubseteq> x)"
0b8e19f588a4 new lemmas about lub
huffman
parents: 39302
diff changeset
    36
  by (simp add: is_lub_below_iff [OF cpo_lubI] is_ub_def)
0b8e19f588a4 new lemmas about lub
huffman
parents: 39302
diff changeset
    37
40500
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40498
diff changeset
    38
lemma lub_below: "\<lbrakk>chain S; \<And>i. S i \<sqsubseteq> x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x"
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40498
diff changeset
    39
  by (simp add: lub_below_iff)
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40498
diff changeset
    40
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40498
diff changeset
    41
lemma below_lub: "\<lbrakk>chain S; x \<sqsubseteq> S i\<rbrakk> \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. S i)"
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40498
diff changeset
    42
  by (erule below_trans, erule is_ub_thelub)
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40498
diff changeset
    43
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    44
lemma lub_range_mono: "\<lbrakk>range X \<subseteq> range Y; chain Y; chain X\<rbrakk> \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    45
  apply (erule lub_below)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    46
  apply (subgoal_tac "\<exists>j. X i = Y j")
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    47
   apply clarsimp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    48
   apply (erule is_ub_thelub)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    49
  apply auto
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    50
  done
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    51
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    52
lemma lub_range_shift: "chain Y \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    53
  apply (rule below_antisym)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    54
   apply (rule lub_range_mono)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    55
     apply fast
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    56
    apply assumption
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    57
   apply (erule chain_shift)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    58
  apply (rule lub_below)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    59
   apply assumption
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    60
  apply (rule_tac i="i" in below_lub)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    61
   apply (erule chain_shift)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    62
  apply (erule chain_mono)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    63
  apply (rule le_add1)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    64
  done
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    65
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    66
lemma maxinch_is_thelub: "chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = Y i)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    67
  apply (rule iffI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    68
   apply (fast intro!: lub_eqI lub_finch1)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    69
  apply (unfold max_in_chain_def)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    70
  apply (safe intro!: below_antisym)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    71
   apply (fast elim!: chain_mono)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    72
  apply (drule sym)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    73
  apply (force elim!: is_ub_thelub)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    74
  done
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    75
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    76
text \<open>the \<open>\<sqsubseteq>\<close> relation between two chains is preserved by their lubs\<close>
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    77
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    78
lemma lub_mono: "\<lbrakk>chain X; chain Y; \<And>i. X i \<sqsubseteq> Y i\<rbrakk> \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    79
  by (fast elim: lub_below below_lub)
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    80
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    81
text \<open>the = relation between two chains is preserved by their lubs\<close>
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
    82
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    83
lemma lub_eq: "(\<And>i. X i = Y i) \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)"
35492
5d200f2d7a4f add lemma lub_eq
huffman
parents: 33523
diff changeset
    84
  by simp
5d200f2d7a4f add lemma lub_eq
huffman
parents: 33523
diff changeset
    85
16203
b3268fe39838 added theorem ch2ch_lub
huffman
parents: 16201
diff changeset
    86
lemma ch2ch_lub:
b3268fe39838 added theorem ch2ch_lub
huffman
parents: 16201
diff changeset
    87
  assumes 1: "\<And>j. chain (\<lambda>i. Y i j)"
b3268fe39838 added theorem ch2ch_lub
huffman
parents: 16201
diff changeset
    88
  assumes 2: "\<And>i. chain (\<lambda>j. Y i j)"
b3268fe39838 added theorem ch2ch_lub
huffman
parents: 16201
diff changeset
    89
  shows "chain (\<lambda>i. \<Squnion>j. Y i j)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    90
  apply (rule chainI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    91
  apply (rule lub_mono [OF 2 2])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    92
  apply (rule chainE [OF 1])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    93
  done
16203
b3268fe39838 added theorem ch2ch_lub
huffman
parents: 16201
diff changeset
    94
16201
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
    95
lemma diag_lub:
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
    96
  assumes 1: "\<And>j. chain (\<lambda>i. Y i j)"
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
    97
  assumes 2: "\<And>i. chain (\<lambda>j. Y i j)"
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
    98
  shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>i. Y i i)"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31071
diff changeset
    99
proof (rule below_antisym)
16201
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   100
  have 3: "chain (\<lambda>i. Y i i)"
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   101
    apply (rule chainI)
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31071
diff changeset
   102
    apply (rule below_trans)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   103
     apply (rule chainE [OF 1])
16201
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   104
    apply (rule chainE [OF 2])
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   105
    done
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   106
  have 4: "chain (\<lambda>i. \<Squnion>j. Y i j)"
16203
b3268fe39838 added theorem ch2ch_lub
huffman
parents: 16201
diff changeset
   107
    by (rule ch2ch_lub [OF 1 2])
16201
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   108
  show "(\<Squnion>i. \<Squnion>j. Y i j) \<sqsubseteq> (\<Squnion>i. Y i i)"
40500
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40498
diff changeset
   109
    apply (rule lub_below [OF 4])
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40498
diff changeset
   110
    apply (rule lub_below [OF 2])
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40498
diff changeset
   111
    apply (rule below_lub [OF 3])
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31071
diff changeset
   112
    apply (rule below_trans)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   113
     apply (rule chain_mono [OF 1 max.cobounded1])
54863
82acc20ded73 prefer more canonical names for lemmas on min/max
haftmann
parents: 42151
diff changeset
   114
    apply (rule chain_mono [OF 2 max.cobounded2])
16201
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   115
    done
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   116
  show "(\<Squnion>i. Y i i) \<sqsubseteq> (\<Squnion>i. \<Squnion>j. Y i j)"
25923
5fe4b543512e convert lemma lub_mono to rule_format
huffman
parents: 25922
diff changeset
   117
    apply (rule lub_mono [OF 3 4])
16201
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   118
    apply (rule is_ub_thelub [OF 2])
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   119
    done
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   120
qed
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   121
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   122
lemma ex_lub:
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   123
  assumes 1: "\<And>j. chain (\<lambda>i. Y i j)"
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   124
  assumes 2: "\<And>i. chain (\<lambda>j. Y i j)"
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   125
  shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>j. \<Squnion>i. Y i j)"
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   126
  by (simp add: diag_lub 1 2)
16201
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   127
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   128
end
16201
7bb51c8196cb added theorems diag_lub and ex_lub
huffman
parents: 16070
diff changeset
   129
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   130
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   131
subsection \<open>Pointed cpos\<close>
15588
14e3228f18cc arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   132
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   133
text \<open>The class pcpo of pointed cpos\<close>
15588
14e3228f18cc arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   134
29614
1f7b1b0df292 simplified handling of base sort, dropped axclass
haftmann
parents: 29138
diff changeset
   135
class pcpo = cpo +
1f7b1b0df292 simplified handling of base sort, dropped axclass
haftmann
parents: 29138
diff changeset
   136
  assumes least: "\<exists>x. \<forall>y. x \<sqsubseteq> y"
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   137
begin
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   138
61998
b66d2ca1f907 clarified print modes;
wenzelm
parents: 61169
diff changeset
   139
definition bottom :: "'a"  ("\<bottom>")
41429
cf5f025bc3c7 renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents: 40774
diff changeset
   140
  where "bottom = (THE x. \<forall>y. x \<sqsubseteq> y)"
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   141
41429
cf5f025bc3c7 renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents: 40774
diff changeset
   142
lemma minimal [iff]: "\<bottom> \<sqsubseteq> x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   143
  unfolding bottom_def
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   144
  apply (rule the1I2)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   145
   apply (rule ex_ex1I)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   146
    apply (rule least)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   147
   apply (blast intro: below_antisym)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   148
  apply simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   149
  done
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   150
41429
cf5f025bc3c7 renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents: 40774
diff changeset
   151
end
cf5f025bc3c7 renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents: 40774
diff changeset
   152
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   153
text \<open>Old "UU" syntax:\<close>
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   154
41429
cf5f025bc3c7 renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents: 40774
diff changeset
   155
syntax UU :: logic
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   156
translations "UU" \<rightharpoonup> "CONST bottom"
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   157
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68369
diff changeset
   158
text \<open>Simproc to rewrite \<^term>\<open>\<bottom> = x\<close> to \<^term>\<open>x = \<bottom>\<close>.\<close>
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   159
setup \<open>Reorient_Proc.add (fn Const(\<^const_name>\<open>bottom\<close>, _) => true | _ => false)\<close>
33523
96730ad673be modernized structure Reorient_Proc;
wenzelm
parents: 31076
diff changeset
   160
simproc_setup reorient_bottom ("\<bottom> = x") = Reorient_Proc.proc
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   161
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68369
diff changeset
   162
text \<open>useful lemmas about \<^term>\<open>\<bottom>\<close>\<close>
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   163
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   164
lemma below_bottom_iff [simp]: "x \<sqsubseteq> \<bottom> \<longleftrightarrow> x = \<bottom>"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   165
  by (simp add: po_eq_conv)
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   166
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   167
lemma eq_bottom_iff: "x = \<bottom> \<longleftrightarrow> x \<sqsubseteq> \<bottom>"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   168
  by simp
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   169
41430
1aa23e9f2c87 change some lemma names containing 'UU' to 'bottom'
huffman
parents: 41429
diff changeset
   170
lemma bottomI: "x \<sqsubseteq> \<bottom> \<Longrightarrow> x = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   171
  by (subst eq_bottom_iff)
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   172
40045
e0f372e18f3e add lemma lub_eq_bottom_iff
huffman
parents: 39969
diff changeset
   173
lemma lub_eq_bottom_iff: "chain Y \<Longrightarrow> (\<Squnion>i. Y i) = \<bottom> \<longleftrightarrow> (\<forall>i. Y i = \<bottom>)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   174
  by (simp only: eq_bottom_iff lub_below_iff)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   175
40045
e0f372e18f3e add lemma lub_eq_bottom_iff
huffman
parents: 39969
diff changeset
   176
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   177
subsection \<open>Chain-finite and flat cpos\<close>
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
   178
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   179
text \<open>further useful classes for HOLCF domains\<close>
15588
14e3228f18cc arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   180
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   181
class chfin = po +
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   182
  assumes chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n Y"
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   183
begin
25814
eb181909e7a4 new axclass finite_po < finite, po
huffman
parents: 25781
diff changeset
   184
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   185
subclass cpo
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   186
  apply standard
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   187
  apply (frule chfin)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   188
  apply (blast intro: lub_finch1)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   189
  done
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
   190
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   191
lemma chfin2finch: "chain Y \<Longrightarrow> finite_chain Y"
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   192
  by (simp add: chfin finite_chain_def)
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   193
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   194
end
15588
14e3228f18cc arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   195
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   196
class flat = pcpo +
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   197
  assumes ax_flat: "x \<sqsubseteq> y \<Longrightarrow> x = \<bottom> \<or> x = y"
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   198
begin
15640
2d1d6ea579a1 chfin now a subclass of po, proved instance chfin < cpo
huffman
parents: 15588
diff changeset
   199
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   200
subclass chfin
68369
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   201
proof
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   202
  fix Y
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   203
  assume *: "chain Y"
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   204
  show "\<exists>n. max_in_chain n Y"
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   205
    apply (unfold max_in_chain_def)
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   206
    apply (cases "\<forall>i. Y i = \<bottom>")
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   207
     apply simp
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   208
    apply simp
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   209
    apply (erule exE)
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   210
    apply (rule_tac x="i" in exI)
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   211
    apply clarify
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   212
    using * apply (blast dest: chain_mono ax_flat)
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   213
    done
6989752bba4b tuned proofs;
wenzelm
parents: 67312
diff changeset
   214
qed
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
   215
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   216
lemma flat_below_iff: "x \<sqsubseteq> y \<longleftrightarrow> x = \<bottom> \<or> x = y"
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   217
  by (safe dest!: ax_flat)
25826
f9aa43287e42 new lemma flat_less_iff
huffman
parents: 25814
diff changeset
   218
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   219
lemma flat_eq: "a \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)"
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   220
  by (safe dest!: ax_flat)
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
   221
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   222
end
15563
9e125b675253 converted to new-style theory
huffman
parents: 14981
diff changeset
   223
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   224
subsection \<open>Discrete cpos\<close>
26023
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   225
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31071
diff changeset
   226
class discrete_cpo = below +
29614
1f7b1b0df292 simplified handling of base sort, dropped axclass
haftmann
parents: 29138
diff changeset
   227
  assumes discrete_cpo [simp]: "x \<sqsubseteq> y \<longleftrightarrow> x = y"
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   228
begin
26023
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   229
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   230
subclass po
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   231
  by standard simp_all
26023
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   232
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   233
text \<open>In a discrete cpo, every chain is constant\<close>
26023
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   234
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   235
lemma discrete_chain_const:
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   236
  assumes S: "chain S"
26023
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   237
  shows "\<exists>x. S = (\<lambda>i. x)"
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   238
proof (intro exI ext)
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   239
  fix i :: nat
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   240
  from S le0 have "S 0 \<sqsubseteq> S i" by (rule chain_mono)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   241
  then have "S 0 = S i" by simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   242
  then show "S i = S 0" by (rule sym)
26023
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   243
qed
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   244
40091
1ca61fbd8a79 make discrete_cpo a subclass of chfin; remove chfin instances for fun, cfun
huffman
parents: 40089
diff changeset
   245
subclass chfin
26023
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   246
proof
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   247
  fix S :: "nat \<Rightarrow> 'a"
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   248
  assume S: "chain S"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   249
  then have "\<exists>x. S = (\<lambda>i. x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   250
    by (rule discrete_chain_const)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   251
  then have "max_in_chain 0 S"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   252
    by (auto simp: max_in_chain_def)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   253
  then show "\<exists>i. max_in_chain i S" ..
26023
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   254
qed
29c1e3e98276 new discrete_cpo axclass
huffman
parents: 25923
diff changeset
   255
31071
845a6acd3bf3 localized (complete) partial order classes
haftmann
parents: 31024
diff changeset
   256
end
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 15563
diff changeset
   257
16626
d28314d2dce3 cleaned up
huffman
parents: 16203
diff changeset
   258
end