src/HOL/HOLCF/Product_Cpo.thy
author paulson <lp15@cam.ac.uk>
Wed, 24 Apr 2024 09:21:44 +0100
changeset 80148 b156869b826a
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
Another Nominal example
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41430
diff changeset
     1
(*  Title:      HOL/HOLCF/Product_Cpo.thy
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
     2
    Author:     Franz Regensburger
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
     3
*)
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
     4
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
     5
section \<open>The cpo of cartesian products\<close>
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
     6
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
     7
theory Product_Cpo
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
     8
  imports Adm
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
     9
begin
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    10
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35919
diff changeset
    11
default_sort cpo
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    12
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    13
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
    14
subsection \<open>Unit type is a pcpo\<close>
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    15
40090
d57357cda8bb direct instantiation unit :: discrete_cpo
huffman
parents: 40089
diff changeset
    16
instantiation unit :: discrete_cpo
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    17
begin
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    18
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    19
definition below_unit_def [simp]: "x \<sqsubseteq> (y::unit) \<longleftrightarrow> True"
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    20
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    21
instance
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    22
  by standard simp
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    23
40090
d57357cda8bb direct instantiation unit :: discrete_cpo
huffman
parents: 40089
diff changeset
    24
end
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    25
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    26
instance unit :: pcpo
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    27
  by standard simp
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    28
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    29
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
    30
subsection \<open>Product type is a partial order\<close>
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    31
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37079
diff changeset
    32
instantiation prod :: (below, below) below
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    33
begin
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    34
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67312
diff changeset
    35
definition below_prod_def: "(\<sqsubseteq>) \<equiv> \<lambda>p1 p2. (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)"
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    36
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    37
instance ..
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    38
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    39
end
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    40
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37079
diff changeset
    41
instance prod :: (po, po) po
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    42
proof
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    43
  fix x :: "'a \<times> 'b"
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    44
  show "x \<sqsubseteq> x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    45
    by (simp add: below_prod_def)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    46
next
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    47
  fix x y :: "'a \<times> 'b"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    48
  assume "x \<sqsubseteq> y" "y \<sqsubseteq> x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    49
  then show "x = y"
44066
d74182c93f04 rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
huffman
parents: 42151
diff changeset
    50
    unfolding below_prod_def prod_eq_iff
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
    51
    by (fast intro: below_antisym)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    52
next
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    53
  fix x y z :: "'a \<times> 'b"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    54
  assume "x \<sqsubseteq> y" "y \<sqsubseteq> z"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    55
  then show "x \<sqsubseteq> z"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
    56
    unfolding below_prod_def
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
    57
    by (fast intro: below_trans)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    58
qed
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    59
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    60
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
    61
subsection \<open>Monotonicity of \emph{Pair}, \emph{fst}, \emph{snd}\<close>
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    62
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    63
lemma prod_belowI: "fst p \<sqsubseteq> fst q \<Longrightarrow> snd p \<sqsubseteq> snd q \<Longrightarrow> p \<sqsubseteq> q"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    64
  by (simp add: below_prod_def)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    65
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
    66
lemma Pair_below_iff [simp]: "(a, b) \<sqsubseteq> (c, d) \<longleftrightarrow> a \<sqsubseteq> c \<and> b \<sqsubseteq> d"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    67
  by (simp add: below_prod_def)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    68
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
    69
text \<open>Pair \<open>(_,_)\<close>  is monotone in both arguments\<close>
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    70
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    71
lemma monofun_pair1: "monofun (\<lambda>x. (x, y))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    72
  by (simp add: monofun_def)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    73
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    74
lemma monofun_pair2: "monofun (\<lambda>y. (x, y))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    75
  by (simp add: monofun_def)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    76
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    77
lemma monofun_pair: "x1 \<sqsubseteq> x2 \<Longrightarrow> y1 \<sqsubseteq> y2 \<Longrightarrow> (x1, y1) \<sqsubseteq> (x2, y2)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    78
  by simp
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    79
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    80
lemma ch2ch_Pair [simp]: "chain X \<Longrightarrow> chain Y \<Longrightarrow> chain (\<lambda>i. (X i, Y i))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    81
  by (rule chainI, simp add: chainE)
31112
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
    82
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
    83
text \<open>\<^term>\<open>fst\<close> and \<^term>\<open>snd\<close> are monotone\<close>
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    84
35919
676c6005ad03 add lemmas fst_monofun, snd_monofun
huffman
parents: 35914
diff changeset
    85
lemma fst_monofun: "x \<sqsubseteq> y \<Longrightarrow> fst x \<sqsubseteq> fst y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    86
  by (simp add: below_prod_def)
35919
676c6005ad03 add lemmas fst_monofun, snd_monofun
huffman
parents: 35914
diff changeset
    87
676c6005ad03 add lemmas fst_monofun, snd_monofun
huffman
parents: 35914
diff changeset
    88
lemma snd_monofun: "x \<sqsubseteq> y \<Longrightarrow> snd x \<sqsubseteq> snd y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    89
  by (simp add: below_prod_def)
35919
676c6005ad03 add lemmas fst_monofun, snd_monofun
huffman
parents: 35914
diff changeset
    90
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    91
lemma monofun_fst: "monofun fst"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    92
  by (simp add: monofun_def below_prod_def)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    93
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    94
lemma monofun_snd: "monofun snd"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    95
  by (simp add: monofun_def below_prod_def)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
    96
31112
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
    97
lemmas ch2ch_fst [simp] = ch2ch_monofun [OF monofun_fst]
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
    98
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
    99
lemmas ch2ch_snd [simp] = ch2ch_monofun [OF monofun_snd]
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
   100
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
   101
lemma prod_chain_cases:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   102
  assumes chain: "chain Y"
31112
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
   103
  obtains A B
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
   104
  where "chain A" and "chain B" and "Y = (\<lambda>i. (A i, B i))"
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
   105
proof
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   106
  from chain show "chain (\<lambda>i. fst (Y i))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   107
    by (rule ch2ch_fst)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   108
  from chain show "chain (\<lambda>i. snd (Y i))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   109
    by (rule ch2ch_snd)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   110
  show "Y = (\<lambda>i. (fst (Y i), snd (Y i)))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   111
    by simp
31112
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
   112
qed
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
   113
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   114
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   115
subsection \<open>Product type is a cpo\<close>
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   116
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   117
lemma is_lub_Pair: "range A <<| x \<Longrightarrow> range B <<| y \<Longrightarrow> range (\<lambda>i. (A i, B i)) <<| (x, y)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   118
  by (simp add: is_lub_def is_ub_def below_prod_def)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   119
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   120
lemma lub_Pair: "chain A \<Longrightarrow> chain B \<Longrightarrow> (\<Squnion>i. (A i, B i)) = (\<Squnion>i. A i, \<Squnion>i. B i)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   121
  for A :: "nat \<Rightarrow> 'a::cpo" and B :: "nat \<Rightarrow> 'b::cpo"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   122
  by (fast intro: lub_eqI is_lub_Pair elim: thelubE)
31112
4dcda8ca5d59 new lemmas
huffman
parents: 31076
diff changeset
   123
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40619
diff changeset
   124
lemma is_lub_prod:
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   125
  fixes S :: "nat \<Rightarrow> ('a::cpo \<times> 'b::cpo)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   126
  assumes "chain S"
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   127
  shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   128
  using assms by (auto elim: prod_chain_cases simp: is_lub_Pair cpo_lubI)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   129
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   130
lemma lub_prod: "chain S \<Longrightarrow> (\<Squnion>i. S i) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   131
  for S :: "nat \<Rightarrow> 'a::cpo \<times> 'b::cpo"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   132
  by (rule is_lub_prod [THEN lub_eqI])
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   133
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37079
diff changeset
   134
instance prod :: (cpo, cpo) cpo
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   135
proof
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   136
  fix S :: "nat \<Rightarrow> ('a \<times> 'b)"
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   137
  assume "chain S"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   138
  then have "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40619
diff changeset
   139
    by (rule is_lub_prod)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   140
  then show "\<exists>x. range S <<| x" ..
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   141
qed
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   142
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37079
diff changeset
   143
instance prod :: (discrete_cpo, discrete_cpo) discrete_cpo
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   144
proof
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   145
  fix x y :: "'a \<times> 'b"
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   146
  show "x \<sqsubseteq> y \<longleftrightarrow> x = y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   147
    by (simp add: below_prod_def prod_eq_iff)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   148
qed
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   149
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   150
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   151
subsection \<open>Product type is pointed\<close>
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   152
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40619
diff changeset
   153
lemma minimal_prod: "(\<bottom>, \<bottom>) \<sqsubseteq> p"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   154
  by (simp add: below_prod_def)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   155
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37079
diff changeset
   156
instance prod :: (pcpo, pcpo) pcpo
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   157
  by intro_classes (fast intro: minimal_prod)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   158
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40619
diff changeset
   159
lemma inst_prod_pcpo: "\<bottom> = (\<bottom>, \<bottom>)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   160
  by (rule minimal_prod [THEN bottomI, symmetric])
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   161
40321
d065b195ec89 rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
huffman
parents: 40090
diff changeset
   162
lemma Pair_bottom_iff [simp]: "(x, y) = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   163
  by (simp add: inst_prod_pcpo)
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   164
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   165
lemma fst_strict [simp]: "fst \<bottom> = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   166
  unfolding inst_prod_pcpo by (rule fst_conv)
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   167
33506
afb577487b15 fix name of lemma snd_strict
huffman
parents: 31112
diff changeset
   168
lemma snd_strict [simp]: "snd \<bottom> = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   169
  unfolding inst_prod_pcpo by (rule snd_conv)
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   170
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   171
lemma Pair_strict [simp]: "(\<bottom>, \<bottom>) = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   172
  by simp
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   173
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 58880
diff changeset
   174
lemma split_strict [simp]: "case_prod f \<bottom> = f \<bottom> \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   175
  by (simp add: split_def)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   176
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   177
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   178
subsection \<open>Continuity of \emph{Pair}, \emph{fst}, \emph{snd}\<close>
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   179
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   180
lemma cont_pair1: "cont (\<lambda>x. (x, y))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   181
  apply (rule contI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   182
  apply (rule is_lub_Pair)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   183
   apply (erule cpo_lubI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   184
  apply (rule is_lub_const)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   185
  done
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   186
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   187
lemma cont_pair2: "cont (\<lambda>y. (x, y))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   188
  apply (rule contI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   189
  apply (rule is_lub_Pair)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   190
   apply (rule is_lub_const)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   191
  apply (erule cpo_lubI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   192
  done
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   193
35914
91a7311177c4 remove contlub predicate
huffman
parents: 35900
diff changeset
   194
lemma cont_fst: "cont fst"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   195
  apply (rule contI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   196
  apply (simp add: lub_prod)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   197
  apply (erule cpo_lubI [OF ch2ch_fst])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   198
  done
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   199
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   200
lemma cont_snd: "cont snd"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   201
  apply (rule contI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   202
  apply (simp add: lub_prod)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   203
  apply (erule cpo_lubI [OF ch2ch_snd])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   204
  done
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   205
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36452
diff changeset
   206
lemma cont2cont_Pair [simp, cont2cont]:
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   207
  assumes f: "cont (\<lambda>x. f x)"
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   208
  assumes g: "cont (\<lambda>x. g x)"
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   209
  shows "cont (\<lambda>x. (f x, g x))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   210
  apply (rule cont_apply [OF f cont_pair1])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   211
  apply (rule cont_apply [OF g cont_pair2])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   212
  apply (rule cont_const)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   213
  done
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   214
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36452
diff changeset
   215
lemmas cont2cont_fst [simp, cont2cont] = cont_compose [OF cont_fst]
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29535
diff changeset
   216
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36452
diff changeset
   217
lemmas cont2cont_snd [simp, cont2cont] = cont_compose [OF cont_snd]
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29535
diff changeset
   218
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 44066
diff changeset
   219
lemma cont2cont_case_prod:
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29535
diff changeset
   220
  assumes f1: "\<And>a b. cont (\<lambda>x. f x a b)"
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29535
diff changeset
   221
  assumes f2: "\<And>x b. cont (\<lambda>a. f x a b)"
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29535
diff changeset
   222
  assumes f3: "\<And>x a. cont (\<lambda>b. f x a b)"
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29535
diff changeset
   223
  assumes g: "cont (\<lambda>x. g x)"
39144
23b1e6759359 add lemma cont2cont_split_simple
huffman
parents: 37678
diff changeset
   224
  shows "cont (\<lambda>x. case g x of (a, b) \<Rightarrow> f x a b)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   225
  unfolding split_def
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   226
  apply (rule cont_apply [OF g])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   227
   apply (rule cont_apply [OF cont_fst f2])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   228
   apply (rule cont_apply [OF cont_snd f3])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   229
   apply (rule cont_const)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   230
  apply (rule f1)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   231
  done
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29535
diff changeset
   232
39808
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39144
diff changeset
   233
lemma prod_contI:
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39144
diff changeset
   234
  assumes f1: "\<And>y. cont (\<lambda>x. f (x, y))"
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39144
diff changeset
   235
  assumes f2: "\<And>x. cont (\<lambda>y. f (x, y))"
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39144
diff changeset
   236
  shows "cont f"
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39144
diff changeset
   237
proof -
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39144
diff changeset
   238
  have "cont (\<lambda>(x, y). f (x, y))"
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 44066
diff changeset
   239
    by (intro cont2cont_case_prod f1 f2 cont2cont)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   240
  then show "cont f"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 58880
diff changeset
   241
    by (simp only: case_prod_eta)
39808
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39144
diff changeset
   242
qed
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39144
diff changeset
   243
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   244
lemma prod_cont_iff: "cont f \<longleftrightarrow> (\<forall>y. cont (\<lambda>x. f (x, y))) \<and> (\<forall>x. cont (\<lambda>y. f (x, y)))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   245
  apply safe
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   246
    apply (erule cont_compose [OF _ cont_pair1])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   247
   apply (erule cont_compose [OF _ cont_pair2])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   248
  apply (simp only: prod_contI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   249
  done
39808
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39144
diff changeset
   250
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 44066
diff changeset
   251
lemma cont2cont_case_prod' [simp, cont2cont]:
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29535
diff changeset
   252
  assumes f: "cont (\<lambda>p. f (fst p) (fst (snd p)) (snd (snd p)))"
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29535
diff changeset
   253
  assumes g: "cont (\<lambda>x. g x)"
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 44066
diff changeset
   254
  shows "cont (\<lambda>x. case_prod (f x) (g x))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   255
  using assms by (simp add: cont2cont_case_prod prod_cont_iff)
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   256
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   257
text \<open>The simple version (due to Joachim Breitner) is needed if
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   258
  either element type of the pair is not a cpo.\<close>
39144
23b1e6759359 add lemma cont2cont_split_simple
huffman
parents: 37678
diff changeset
   259
23b1e6759359 add lemma cont2cont_split_simple
huffman
parents: 37678
diff changeset
   260
lemma cont2cont_split_simple [simp, cont2cont]:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   261
  assumes "\<And>a b. cont (\<lambda>x. f x a b)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   262
  shows "cont (\<lambda>x. case p of (a, b) \<Rightarrow> f x a b)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   263
  using assms by (cases p) auto
39144
23b1e6759359 add lemma cont2cont_split_simple
huffman
parents: 37678
diff changeset
   264
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   265
text \<open>Admissibility of predicates on product types.\<close>
40619
84edf7177d73 add lemma adm_prod_case
huffman
parents: 40321
diff changeset
   266
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 44066
diff changeset
   267
lemma adm_case_prod [simp]:
40619
84edf7177d73 add lemma adm_prod_case
huffman
parents: 40321
diff changeset
   268
  assumes "adm (\<lambda>x. P x (fst (f x)) (snd (f x)))"
84edf7177d73 add lemma adm_prod_case
huffman
parents: 40321
diff changeset
   269
  shows "adm (\<lambda>x. case f x of (a, b) \<Rightarrow> P x a b)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   270
  unfolding case_prod_beta using assms .
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   271
40619
84edf7177d73 add lemma adm_prod_case
huffman
parents: 40321
diff changeset
   272
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   273
subsection \<open>Compactness and chain-finiteness\<close>
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   274
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   275
lemma fst_below_iff: "fst x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (y, snd x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   276
  for x :: "'a \<times> 'b"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   277
  by (simp add: below_prod_def)
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   278
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   279
lemma snd_below_iff: "snd x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (fst x, y)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   280
  for x :: "'a \<times> 'b"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   281
  by (simp add: below_prod_def)
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   282
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   283
lemma compact_fst: "compact x \<Longrightarrow> compact (fst x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   284
  by (rule compactI) (simp add: fst_below_iff)
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   285
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   286
lemma compact_snd: "compact x \<Longrightarrow> compact (snd x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   287
  by (rule compactI) (simp add: snd_below_iff)
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   288
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   289
lemma compact_Pair: "compact x \<Longrightarrow> compact y \<Longrightarrow> compact (x, y)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   290
  by (rule compactI) (simp add: below_prod_def)
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   291
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   292
lemma compact_Pair_iff [simp]: "compact (x, y) \<longleftrightarrow> compact x \<and> compact y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   293
  apply (safe intro!: compact_Pair)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   294
   apply (drule compact_fst, simp)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   295
  apply (drule compact_snd, simp)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   296
  done
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   297
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37079
diff changeset
   298
instance prod :: (chfin, chfin) chfin
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   299
  apply intro_classes
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   300
  apply (erule compact_imp_max_in_chain)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   301
  apply (case_tac "\<Squnion>i. Y i", simp)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   302
  done
29535
08824fad8879 add strictness and compactness lemmas to Product_Cpo.thy
huffman
parents: 29533
diff changeset
   303
29531
2eb29775b0b6 add Product_Cpo.thy
huffman
parents:
diff changeset
   304
end