src/HOL/Library/Product_Order.thy
author wenzelm
Mon, 26 Jun 2023 23:20:32 +0200
changeset 78209 50c5be88ad59
parent 77138 c8597292cd41
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51115
7dbd6832a689 consolidation of library theories on product orders
haftmann
parents: 50573
diff changeset
     1
(*  Title:      HOL/Library/Product_Order.thy
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
     3
*)
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
     4
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
     5
section \<open>Pointwise order on product types\<close>
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
     6
51115
7dbd6832a689 consolidation of library theories on product orders
haftmann
parents: 50573
diff changeset
     7
theory Product_Order
63972
c98d1dd7eba1 Library: fix name Product_plus to Product_Plus
hoelzl
parents: 63561
diff changeset
     8
imports Product_Plus
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
     9
begin
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    10
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    11
subsection \<open>Pointwise ordering\<close>
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    12
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    13
instantiation prod :: (ord, ord) ord
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    14
begin
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    15
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    16
definition
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    17
  "x \<le> y \<longleftrightarrow> fst x \<le> fst y \<and> snd x \<le> snd y"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    18
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    19
definition
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    20
  "(x::'a \<times> 'b) < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    21
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    22
instance ..
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    23
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    24
end
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    25
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    26
lemma fst_mono: "x \<le> y \<Longrightarrow> fst x \<le> fst y"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    27
  unfolding less_eq_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    28
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    29
lemma snd_mono: "x \<le> y \<Longrightarrow> snd x \<le> snd y"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    30
  unfolding less_eq_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    31
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    32
lemma Pair_mono: "x \<le> x' \<Longrightarrow> y \<le> y' \<Longrightarrow> (x, y) \<le> (x', y')"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    33
  unfolding less_eq_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    34
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    35
lemma Pair_le [simp]: "(a, b) \<le> (c, d) \<longleftrightarrow> a \<le> c \<and> b \<le> d"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    36
  unfolding less_eq_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    37
77138
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 69895
diff changeset
    38
lemma atLeastAtMost_prod_eq: "{a..b} = {fst a..fst b} \<times> {snd a..snd b}"
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 69895
diff changeset
    39
  by (auto simp: less_eq_prod_def)
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 69895
diff changeset
    40
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    41
instance prod :: (preorder, preorder) preorder
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    42
proof
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    43
  fix x y z :: "'a \<times> 'b"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    44
  show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    45
    by (rule less_prod_def)
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    46
  show "x \<le> x"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    47
    unfolding less_eq_prod_def
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    48
    by fast
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    49
  assume "x \<le> y" and "y \<le> z" thus "x \<le> z"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    50
    unfolding less_eq_prod_def
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    51
    by (fast elim: order_trans)
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    52
qed
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    53
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    54
instance prod :: (order, order) order
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
    55
  by standard auto
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    56
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    57
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    58
subsection \<open>Binary infimum and supremum\<close>
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    59
54776
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
    60
instantiation prod :: (inf, inf) inf
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    61
begin
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    62
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
    63
definition "inf x y = (inf (fst x) (fst y), inf (snd x) (snd y))"
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    64
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    65
lemma inf_Pair_Pair [simp]: "inf (a, b) (c, d) = (inf a c, inf b d)"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    66
  unfolding inf_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    67
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    68
lemma fst_inf [simp]: "fst (inf x y) = inf (fst x) (fst y)"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    69
  unfolding inf_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    70
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    71
lemma snd_inf [simp]: "snd (inf x y) = inf (snd x) (snd y)"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    72
  unfolding inf_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    73
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
    74
instance ..
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
    75
54776
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
    76
end
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
    77
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
    78
instance prod :: (semilattice_inf, semilattice_inf) semilattice_inf
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
    79
  by standard auto
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    80
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    81
54776
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
    82
instantiation prod :: (sup, sup) sup
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    83
begin
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    84
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    85
definition
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    86
  "sup x y = (sup (fst x) (fst y), sup (snd x) (snd y))"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    87
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    88
lemma sup_Pair_Pair [simp]: "sup (a, b) (c, d) = (sup a c, sup b d)"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    89
  unfolding sup_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    90
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    91
lemma fst_sup [simp]: "fst (sup x y) = sup (fst x) (fst y)"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    92
  unfolding sup_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    93
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    94
lemma snd_sup [simp]: "snd (sup x y) = sup (snd x) (snd y)"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    95
  unfolding sup_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
    96
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
    97
instance ..
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
    98
54776
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
    99
end
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   100
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   101
instance prod :: (semilattice_sup, semilattice_sup) semilattice_sup
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
   102
  by standard auto
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   103
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   104
instance prod :: (lattice, lattice) lattice ..
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   105
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   106
instance prod :: (distrib_lattice, distrib_lattice) distrib_lattice
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
   107
  by standard (auto simp add: sup_inf_distrib1)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   108
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   109
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   110
subsection \<open>Top and bottom elements\<close>
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   111
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   112
instantiation prod :: (top, top) top
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   113
begin
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   114
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   115
definition
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   116
  "top = (top, top)"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   117
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   118
instance ..
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   119
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   120
end
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   121
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   122
lemma fst_top [simp]: "fst top = top"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   123
  unfolding top_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   124
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   125
lemma snd_top [simp]: "snd top = top"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   126
  unfolding top_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   127
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   128
lemma Pair_top_top: "(top, top) = top"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   129
  unfolding top_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   130
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   131
instance prod :: (order_top, order_top) order_top
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
   132
  by standard (auto simp add: top_prod_def)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   133
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   134
instantiation prod :: (bot, bot) bot
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   135
begin
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   136
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   137
definition
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   138
  "bot = (bot, bot)"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   139
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   140
instance ..
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   141
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   142
end
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   143
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   144
lemma fst_bot [simp]: "fst bot = bot"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   145
  unfolding bot_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   146
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   147
lemma snd_bot [simp]: "snd bot = bot"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   148
  unfolding bot_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   149
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   150
lemma Pair_bot_bot: "(bot, bot) = bot"
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   151
  unfolding bot_prod_def by simp
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   152
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   153
instance prod :: (order_bot, order_bot) order_bot
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
   154
  by standard (auto simp add: bot_prod_def)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   155
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   156
instance prod :: (bounded_lattice, bounded_lattice) bounded_lattice ..
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   157
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   158
instance prod :: (boolean_algebra, boolean_algebra) boolean_algebra
62053
1c8252d07e32 tuned proofs;
wenzelm
parents: 61166
diff changeset
   159
  by standard (auto simp add: prod_eqI diff_eq)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   160
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   161
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   162
subsection \<open>Complete lattice operations\<close>
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   163
54776
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   164
instantiation prod :: (Inf, Inf) Inf
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   165
begin
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   166
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67829
diff changeset
   167
definition "Inf A = (INF x\<in>A. fst x, INF x\<in>A. snd x)"
54776
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   168
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
   169
instance ..
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
   170
54776
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   171
end
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   172
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   173
instantiation prod :: (Sup, Sup) Sup
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   174
begin
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   175
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67829
diff changeset
   176
definition "Sup A = (SUP x\<in>A. fst x, SUP x\<in>A. snd x)"
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   177
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
   178
instance ..
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
   179
54776
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   180
end
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   181
54776
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   182
instance prod :: (conditionally_complete_lattice, conditionally_complete_lattice)
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   183
    conditionally_complete_lattice
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
   184
  by standard (force simp: less_eq_prod_def Inf_prod_def Sup_prod_def bdd_below_def bdd_above_def
62053
1c8252d07e32 tuned proofs;
wenzelm
parents: 61166
diff changeset
   185
    intro!: cInf_lower cSup_upper cInf_greatest cSup_least)+
54776
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   186
db890d9fc5c2 ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents: 52729
diff changeset
   187
instance prod :: (complete_lattice, complete_lattice) complete_lattice
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60580
diff changeset
   188
  by standard (simp_all add: less_eq_prod_def Inf_prod_def Sup_prod_def
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 51542
diff changeset
   189
    INF_lower SUP_upper le_INF_iff SUP_le_iff bot_prod_def top_prod_def)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   190
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   191
lemma fst_Inf: "fst (Inf A) = (INF x\<in>A. fst x)"
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   192
  by (simp add: Inf_prod_def)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   193
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   194
lemma fst_INF: "fst (INF x\<in>A. f x) = (INF x\<in>A. fst (f x))"
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   195
  by (simp add: fst_Inf image_image)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   196
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   197
lemma fst_Sup: "fst (Sup A) = (SUP x\<in>A. fst x)"
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   198
  by (simp add: Sup_prod_def)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   199
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67829
diff changeset
   200
lemma fst_SUP: "fst (SUP x\<in>A. f x) = (SUP x\<in>A. fst (f x))"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   201
  by (simp add: fst_Sup image_image)
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   202
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   203
lemma snd_Inf: "snd (Inf A) = (INF x\<in>A. snd x)"
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   204
  by (simp add: Inf_prod_def)
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   205
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   206
lemma snd_INF: "snd (INF x\<in>A. f x) = (INF x\<in>A. snd (f x))"
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   207
  by (simp add: snd_Inf image_image)
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   208
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   209
lemma snd_Sup: "snd (Sup A) = (SUP x\<in>A. snd x)"
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   210
  by (simp add: Sup_prod_def)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   211
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67829
diff changeset
   212
lemma snd_SUP: "snd (SUP x\<in>A. f x) = (SUP x\<in>A. snd (f x))"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   213
  by (simp add: snd_Sup image_image)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   214
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   215
lemma INF_Pair: "(INF x\<in>A. (f x, g x)) = (INF x\<in>A. f x, INF x\<in>A. g x)"
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   216
  by (simp add: Inf_prod_def image_image)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   217
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67829
diff changeset
   218
lemma SUP_Pair: "(SUP x\<in>A. (f x, g x)) = (SUP x\<in>A. f x, SUP x\<in>A. g x)"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   219
  by (simp add: Sup_prod_def image_image)
44006
b9839fad3bb6 new theory HOL/Library/Product_Lattice.thy
huffman
parents:
diff changeset
   220
50535
2464d77527c4 contribution by A. Colgio
nipkow
parents: 45007
diff changeset
   221
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   222
text \<open>Alternative formulations for set infima and suprema over the product
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   223
of two complete lattices:\<close>
50535
2464d77527c4 contribution by A. Colgio
nipkow
parents: 45007
diff changeset
   224
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69861
diff changeset
   225
lemma INF_prod_alt_def: \<^marker>\<open>contributor \<open>Alessandro Coglio\<close>\<close>
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   226
  "Inf (f ` A) = (Inf ((fst \<circ> f) ` A), Inf ((snd \<circ> f) ` A))"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   227
  by (simp add: Inf_prod_def image_image)
50535
2464d77527c4 contribution by A. Colgio
nipkow
parents: 45007
diff changeset
   228
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69861
diff changeset
   229
lemma SUP_prod_alt_def: \<^marker>\<open>contributor \<open>Alessandro Coglio\<close>\<close>
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   230
  "Sup (f ` A) = (Sup ((fst \<circ> f) ` A), Sup((snd \<circ> f) ` A))"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   231
  by (simp add: Sup_prod_def image_image)
50535
2464d77527c4 contribution by A. Colgio
nipkow
parents: 45007
diff changeset
   232
2464d77527c4 contribution by A. Colgio
nipkow
parents: 45007
diff changeset
   233
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   234
subsection \<open>Complete distributive lattices\<close>
50535
2464d77527c4 contribution by A. Colgio
nipkow
parents: 45007
diff changeset
   235
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69861
diff changeset
   236
instance prod :: (complete_distrib_lattice, complete_distrib_lattice) complete_distrib_lattice \<^marker>\<open>contributor \<open>Alessandro Coglio\<close>\<close>
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   237
proof
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   238
  fix A::"('a\<times>'b) set set"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   239
  show "Inf (Sup ` A) \<le> Sup (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69313
diff changeset
   240
    by (simp add: Inf_prod_def Sup_prod_def INF_SUP_set image_image)
50535
2464d77527c4 contribution by A. Colgio
nipkow
parents: 45007
diff changeset
   241
qed
2464d77527c4 contribution by A. Colgio
nipkow
parents: 45007
diff changeset
   242
63561
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   243
subsection \<open>Bekic's Theorem\<close>
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   244
text \<open>
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   245
  Simultaneous fixed points over pairs can be written in terms of separate fixed points.
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   246
  Transliterated from HOLCF.Fix by Peter Gammie
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   247
\<close>
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   248
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   249
lemma lfp_prod:
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   250
  fixes F :: "'a::complete_lattice \<times> 'b::complete_lattice \<Rightarrow> 'a \<times> 'b"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   251
  assumes "mono F"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   252
  shows "lfp F = (lfp (\<lambda>x. fst (F (x, lfp (\<lambda>y. snd (F (x, y)))))),
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   253
                 (lfp (\<lambda>y. snd (F (lfp (\<lambda>x. fst (F (x, lfp (\<lambda>y. snd (F (x, y)))))), y)))))"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   254
  (is "lfp F = (?x, ?y)")
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   255
proof(rule lfp_eqI[OF assms])
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   256
  have 1: "fst (F (?x, ?y)) = ?x"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   257
    by (rule trans [symmetric, OF lfp_unfold])
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   258
       (blast intro!: monoI monoD[OF assms(1)] fst_mono snd_mono Pair_mono lfp_mono)+
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   259
  have 2: "snd (F (?x, ?y)) = ?y"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   260
    by (rule trans [symmetric, OF lfp_unfold])
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   261
       (blast intro!: monoI monoD[OF assms(1)] fst_mono snd_mono Pair_mono lfp_mono)+
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   262
  from 1 2 show "F (?x, ?y) = (?x, ?y)" by (simp add: prod_eq_iff)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   263
next
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   264
  fix z assume F_z: "F z = z"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   265
  obtain x y where z: "z = (x, y)" by (rule prod.exhaust)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   266
  from F_z z have F_x: "fst (F (x, y)) = x" by simp
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   267
  from F_z z have F_y: "snd (F (x, y)) = y" by simp
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   268
  let ?y1 = "lfp (\<lambda>y. snd (F (x, y)))"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   269
  have "?y1 \<le> y" by (rule lfp_lowerbound, simp add: F_y)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   270
  hence "fst (F (x, ?y1)) \<le> fst (F (x, y))"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   271
    by (simp add: assms fst_mono monoD)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   272
  hence "fst (F (x, ?y1)) \<le> x" using F_x by simp
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   273
  hence 1: "?x \<le> x" by (simp add: lfp_lowerbound)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   274
  hence "snd (F (?x, y)) \<le> snd (F (x, y))"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   275
    by (simp add: assms snd_mono monoD)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   276
  hence "snd (F (?x, y)) \<le> y" using F_y by simp
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   277
  hence 2: "?y \<le> y" by (simp add: lfp_lowerbound)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   278
  show "(?x, ?y) \<le> z" using z 1 2 by simp
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   279
qed
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   280
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   281
lemma gfp_prod:
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   282
  fixes F :: "'a::complete_lattice \<times> 'b::complete_lattice \<Rightarrow> 'a \<times> 'b"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   283
  assumes "mono F"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   284
  shows "gfp F = (gfp (\<lambda>x. fst (F (x, gfp (\<lambda>y. snd (F (x, y)))))),
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   285
                 (gfp (\<lambda>y. snd (F (gfp (\<lambda>x. fst (F (x, gfp (\<lambda>y. snd (F (x, y)))))), y)))))"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   286
  (is "gfp F = (?x, ?y)")
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   287
proof(rule gfp_eqI[OF assms])
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   288
  have 1: "fst (F (?x, ?y)) = ?x"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   289
    by (rule trans [symmetric, OF gfp_unfold])
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   290
       (blast intro!: monoI monoD[OF assms(1)] fst_mono snd_mono Pair_mono gfp_mono)+
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   291
  have 2: "snd (F (?x, ?y)) = ?y"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   292
    by (rule trans [symmetric, OF gfp_unfold])
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   293
       (blast intro!: monoI monoD[OF assms(1)] fst_mono snd_mono Pair_mono gfp_mono)+
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   294
  from 1 2 show "F (?x, ?y) = (?x, ?y)" by (simp add: prod_eq_iff)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   295
next
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   296
  fix z assume F_z: "F z = z"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   297
  obtain x y where z: "z = (x, y)" by (rule prod.exhaust)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   298
  from F_z z have F_x: "fst (F (x, y)) = x" by simp
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   299
  from F_z z have F_y: "snd (F (x, y)) = y" by simp
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   300
  let ?y1 = "gfp (\<lambda>y. snd (F (x, y)))"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   301
  have "y \<le> ?y1" by (rule gfp_upperbound, simp add: F_y)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   302
  hence "fst (F (x, y)) \<le> fst (F (x, ?y1))"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   303
    by (simp add: assms fst_mono monoD)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   304
  hence "x \<le> fst (F (x, ?y1))" using F_x by simp
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   305
  hence 1: "x \<le> ?x" by (simp add: gfp_upperbound)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   306
  hence "snd (F (x, y)) \<le> snd (F (?x, y))"
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   307
    by (simp add: assms snd_mono monoD)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   308
  hence "y \<le> snd (F (?x, y))" using F_y by simp
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   309
  hence 2: "y \<le> ?y" by (simp add: gfp_upperbound)
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   310
  show "z \<le> (?x, ?y)" using z 1 2 by simp
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   311
qed
fba08009ff3e add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents: 62343
diff changeset
   312
51115
7dbd6832a689 consolidation of library theories on product orders
haftmann
parents: 50573
diff changeset
   313
end