src/HOL/Library/Complete_Partial_Order2.thy
author wenzelm
Wed, 12 Mar 2025 11:39:00 +0100
changeset 82265 4b875a4c83b0
parent 81982 bd2779a1da2c
permissions -rw-r--r--
update for release;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62858
d72a6f9ee690 tuned headers;
wenzelm
parents: 62837
diff changeset
     1
(*  Title:      HOL/Library/Complete_Partial_Order2.thy
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
     2
    Author:     Andreas Lochbihler, ETH Zurich
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
     3
*)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
     4
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
     5
section \<open>Formalisation of chain-complete partial orders, continuity and admissibility\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
     6
81281
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
     7
theory Complete_Partial_Order2
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
     8
  imports Main
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
     9
begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    10
74334
ead56ad40e15 bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
wenzelm
parents: 73411
diff changeset
    11
unbundle lattice_syntax
ead56ad40e15 bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
wenzelm
parents: 73411
diff changeset
    12
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    13
lemma chain_transfer [transfer_rule]:
63343
fb5d8a50c641 bundle lifting_syntax;
wenzelm
parents: 63243
diff changeset
    14
  includes lifting_syntax
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
    15
  shows "((A ===> A ===> (=)) ===> rel_set A ===> (=)) Complete_Partial_Order.chain Complete_Partial_Order.chain"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    16
  unfolding chain_def[abs_def] by transfer_prover
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    17
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    18
lemma linorder_chain [simp, intro!]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    19
  fixes Y :: "_ :: linorder set"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
    20
  shows "Complete_Partial_Order.chain (\<le>) Y"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    21
  by(auto intro: chainI)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    22
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    23
lemma fun_lub_apply: "\<And>Sup. fun_lub Sup Y x = Sup ((\<lambda>f. f x) ` Y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    24
  by(simp add: fun_lub_def image_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    25
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    26
lemma fun_lub_empty [simp]: "fun_lub lub {} = (\<lambda>_. lub {})"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    27
  by(rule ext)(simp add: fun_lub_apply)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    28
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    29
lemma chain_fun_ordD: 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    30
  assumes "Complete_Partial_Order.chain (fun_ord le) Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    31
  shows "Complete_Partial_Order.chain le ((\<lambda>f. f x) ` Y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    32
  by(rule chainI)(auto dest: chainD[OF assms] simp add: fun_ord_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    33
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    34
lemma chain_Diff:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    35
  "Complete_Partial_Order.chain ord A
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    36
  \<Longrightarrow> Complete_Partial_Order.chain ord (A - B)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    37
  by(erule chain_subset) blast
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    38
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    39
lemma chain_rel_prodD1:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    40
  "Complete_Partial_Order.chain (rel_prod orda ordb) Y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    41
  \<Longrightarrow> Complete_Partial_Order.chain orda (fst ` Y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    42
  by(auto 4 3 simp add: chain_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    43
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    44
lemma chain_rel_prodD2:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    45
  "Complete_Partial_Order.chain (rel_prod orda ordb) Y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    46
  \<Longrightarrow> Complete_Partial_Order.chain ordb (snd ` Y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    47
  by(auto 4 3 simp add: chain_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    48
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    49
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    50
context ccpo begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    51
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
    52
lemma ccpo_fun: "class.ccpo (fun_lub Sup) (fun_ord (\<le>)) (mk_less (fun_ord (\<le>)))"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    53
  by standard (auto 4 3 simp add: mk_less_def fun_ord_def fun_lub_apply
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    54
      intro: order.trans order.antisym chain_imageI ccpo_Sup_upper ccpo_Sup_least)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    55
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
    56
lemma ccpo_Sup_below_iff: "Complete_Partial_Order.chain (\<le>) Y \<Longrightarrow> Sup Y \<le> x \<longleftrightarrow> (\<forall>y\<in>Y. y \<le> x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
    57
  by (meson local.ccpo_Sup_least local.ccpo_Sup_upper local.dual_order.trans)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    58
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    59
lemma Sup_minus_bot: 
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
    60
  assumes chain: "Complete_Partial_Order.chain (\<le>) A"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    61
  shows "\<Squnion>(A - {\<Squnion>{}}) = \<Squnion>A"
63649
e690d6f2185b tuned proofs;
wenzelm
parents: 63343
diff changeset
    62
    (is "?lhs = ?rhs")
73411
1f1366966296 avoid name clash
haftmann
parents: 70961
diff changeset
    63
proof (rule order.antisym)
63649
e690d6f2185b tuned proofs;
wenzelm
parents: 63343
diff changeset
    64
  show "?lhs \<le> ?rhs"
e690d6f2185b tuned proofs;
wenzelm
parents: 63343
diff changeset
    65
    by (blast intro: ccpo_Sup_least chain_Diff[OF chain] ccpo_Sup_upper[OF chain])
e690d6f2185b tuned proofs;
wenzelm
parents: 63343
diff changeset
    66
  show "?rhs \<le> ?lhs"
e690d6f2185b tuned proofs;
wenzelm
parents: 63343
diff changeset
    67
  proof (rule ccpo_Sup_least [OF chain])
e690d6f2185b tuned proofs;
wenzelm
parents: 63343
diff changeset
    68
    show "x \<in> A \<Longrightarrow> x \<le> ?lhs" for x
e690d6f2185b tuned proofs;
wenzelm
parents: 63343
diff changeset
    69
      by (cases "x = \<Squnion>{}")
e690d6f2185b tuned proofs;
wenzelm
parents: 63343
diff changeset
    70
        (blast intro: ccpo_Sup_least chain_empty ccpo_Sup_upper[OF chain_Diff[OF chain]])+
e690d6f2185b tuned proofs;
wenzelm
parents: 63343
diff changeset
    71
  qed
e690d6f2185b tuned proofs;
wenzelm
parents: 63343
diff changeset
    72
qed
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    73
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    74
lemma mono_lub:
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 76749
diff changeset
    75
  fixes le_b (infix \<open>\<sqsubseteq>\<close> 60)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
    76
  assumes chain: "Complete_Partial_Order.chain (fun_ord (\<le>)) Y"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
    77
  and mono: "\<And>f. f \<in> Y \<Longrightarrow> monotone le_b (\<le>) f"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
    78
  shows "monotone (\<sqsubseteq>) (\<le>) (fun_lub Sup Y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    79
proof(rule monotoneI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    80
  fix x y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    81
  assume "x \<sqsubseteq> y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    82
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
    83
  have chain'': "\<And>x. Complete_Partial_Order.chain (\<le>) ((\<lambda>f. f x) ` Y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    84
    using chain by(rule chain_imageI)(simp add: fun_ord_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    85
  then show "fun_lub Sup Y x \<le> fun_lub Sup Y y" unfolding fun_lub_apply
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    86
  proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    87
    fix x'
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    88
    assume "x' \<in> (\<lambda>f. f x) ` Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    89
    then obtain f where "f \<in> Y" "x' = f x" by blast
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
    90
    note \<open>x' = f x\<close> also
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
    91
    from \<open>f \<in> Y\<close> \<open>x \<sqsubseteq> y\<close> have "f x \<le> f y" by(blast dest: mono monotoneD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    92
    also have "\<dots> \<le> \<Squnion>((\<lambda>f. f y) ` Y)" using chain''
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
    93
      by(rule ccpo_Sup_upper)(simp add: \<open>f \<in> Y\<close>)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    94
    finally show "x' \<le> \<Squnion>((\<lambda>f. f y) ` Y)" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    95
  qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    96
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    97
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
    98
context
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 76749
diff changeset
    99
  fixes le_b (infix \<open>\<sqsubseteq>\<close> 60) and Y f
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   100
  assumes chain: "Complete_Partial_Order.chain le_b Y" 
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   101
  and mono1: "\<And>y. y \<in> Y \<Longrightarrow> monotone le_b (\<le>) (\<lambda>x. f x y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   102
  and mono2: "\<And>x a b. \<lbrakk> x \<in> Y; a \<sqsubseteq> b; a \<in> Y; b \<in> Y \<rbrakk> \<Longrightarrow> f x a \<le> f x b"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   103
begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   104
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   105
lemma Sup_mono: 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   106
  assumes le: "x \<sqsubseteq> y" and x: "x \<in> Y" and y: "y \<in> Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   107
  shows "\<Squnion>(f x ` Y) \<le> \<Squnion>(f y ` Y)" (is "_ \<le> ?rhs")
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   108
proof(rule ccpo_Sup_least)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   109
  from chain show chain': "Complete_Partial_Order.chain (\<le>) (f x ` Y)" when "x \<in> Y" for x
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   110
    by(rule chain_imageI) (insert that, auto dest: mono2)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   111
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   112
  fix x'
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   113
  assume "x' \<in> f x ` Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   114
  then obtain y' where "y' \<in> Y" "x' = f x y'" by blast note this(2)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   115
  also from mono1[OF \<open>y' \<in> Y\<close>] le have "\<dots> \<le> f y y'" by(rule monotoneD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   116
  also have "\<dots> \<le> ?rhs" using chain'[OF y]
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   117
    by (auto intro!: ccpo_Sup_upper simp add: \<open>y' \<in> Y\<close>)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   118
  finally show "x' \<le> ?rhs" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   119
qed(rule x)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   120
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   121
lemma diag_Sup: "\<Squnion>((\<lambda>x. \<Squnion>(f x ` Y)) ` Y) = \<Squnion>((\<lambda>x. f x x) ` Y)" (is "?lhs = ?rhs")
73411
1f1366966296 avoid name clash
haftmann
parents: 70961
diff changeset
   122
proof(rule order.antisym)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   123
  have chain1: "Complete_Partial_Order.chain (\<le>) ((\<lambda>x. \<Squnion>(f x ` Y)) ` Y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   124
    using chain by(rule chain_imageI)(rule Sup_mono)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   125
  have chain2: "\<And>y'. y' \<in> Y \<Longrightarrow> Complete_Partial_Order.chain (\<le>) (f y' ` Y)" using chain
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   126
    by(rule chain_imageI)(auto dest: mono2)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   127
  have chain3: "Complete_Partial_Order.chain (\<le>) ((\<lambda>x. f x x) ` Y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   128
    using chain by(rule chain_imageI)(auto intro: monotoneD[OF mono1] mono2 order.trans)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   129
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   130
  show "?lhs \<le> ?rhs" using chain1
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   131
  proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   132
    fix x'
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   133
    assume "x' \<in> (\<lambda>x. \<Squnion>(f x ` Y)) ` Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   134
    then obtain y' where "y' \<in> Y" "x' = \<Squnion>(f y' ` Y)" by blast note this(2)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   135
    also have "\<dots> \<le> ?rhs" using chain2[OF \<open>y' \<in> Y\<close>]
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   136
    proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   137
      fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   138
      assume "x \<in> f y' ` Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   139
      then obtain y where "y \<in> Y" and x: "x = f y' y" by blast
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62858
diff changeset
   140
      define y'' where "y'' = (if y \<sqsubseteq> y' then y' else y)"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   141
      from chain \<open>y \<in> Y\<close> \<open>y' \<in> Y\<close> have "y \<sqsubseteq> y' \<or> y' \<sqsubseteq> y" by(rule chainD)
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   142
      hence "f y' y \<le> f y'' y''" using \<open>y \<in> Y\<close> \<open>y' \<in> Y\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   143
        by(auto simp add: y''_def intro: mono2 monotoneD[OF mono1])
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   144
      also from \<open>y \<in> Y\<close> \<open>y' \<in> Y\<close> have "y'' \<in> Y" by(simp add: y''_def)
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   145
      from chain3 have "f y'' y'' \<le> ?rhs" by(rule ccpo_Sup_upper)(simp add: \<open>y'' \<in> Y\<close>)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   146
      finally show "x \<le> ?rhs" by(simp add: x)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   147
    qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   148
    finally show "x' \<le> ?rhs" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   149
  qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   150
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   151
  show "?rhs \<le> ?lhs" using chain3
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   152
  proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   153
    fix y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   154
    assume "y \<in> (\<lambda>x. f x x) ` Y"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   155
    then obtain x where "x \<in> Y" and "y = f x x" by blast
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   156
    then show "y \<le> ?lhs"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   157
      by (metis (no_types, lifting) chain1 chain2 imageI ccpo_Sup_upper order.trans)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   158
  qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   159
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   160
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   161
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   162
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   163
lemma Sup_image_mono_le:
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 76749
diff changeset
   164
  fixes le_b (infix \<open>\<sqsubseteq>\<close> 60) and Sup_b (\<open>\<Or>\<close>)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   165
  assumes ccpo: "class.ccpo Sup_b (\<sqsubseteq>) lt_b"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   166
  assumes chain: "Complete_Partial_Order.chain (\<sqsubseteq>) Y"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   167
  and mono: "\<And>x y. \<lbrakk> x \<sqsubseteq> y; x \<in> Y \<rbrakk> \<Longrightarrow> f x \<le> f y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   168
  shows "Sup (f ` Y) \<le> f (\<Or>Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   169
proof(rule ccpo_Sup_least)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   170
  show "Complete_Partial_Order.chain (\<le>) (f ` Y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   171
    using chain by(rule chain_imageI)(rule mono)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   172
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   173
  fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   174
  assume "x \<in> f ` Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   175
  then obtain y where "y \<in> Y" and "x = f y" by blast note this(2)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   176
  also have "y \<sqsubseteq> \<Or>Y" using ccpo chain \<open>y \<in> Y\<close> by(rule ccpo.ccpo_Sup_upper)
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   177
  hence "f y \<le> f (\<Or>Y)" using \<open>y \<in> Y\<close> by(rule mono)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   178
  finally show "x \<le> \<dots>" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   179
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   180
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   181
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   182
lemma swap_Sup:
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 76749
diff changeset
   183
  fixes le_b (infix \<open>\<sqsubseteq>\<close> 60)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   184
  assumes Y: "Complete_Partial_Order.chain (\<sqsubseteq>) Y"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   185
  and Z: "Complete_Partial_Order.chain (fun_ord (\<le>)) Z"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   186
  and mono: "\<And>f. f \<in> Z \<Longrightarrow> monotone (\<sqsubseteq>) (\<le>) f"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   187
  shows "\<Squnion>((\<lambda>x. \<Squnion>(x ` Y)) ` Z) = \<Squnion>((\<lambda>x. \<Squnion>((\<lambda>f. f x) ` Z)) ` Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   188
  (is "?lhs = ?rhs")
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   189
proof(cases "Y = {}")
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   190
  case True
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   191
  then show ?thesis
69546
27dae626822b prefer naming convention from datatype package for strong congruence rules
haftmann
parents: 69164
diff changeset
   192
    by (simp add: image_constant_conv cong del: SUP_cong_simp)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   193
next
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   194
  case False
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   195
  have chain1: "\<And>f. f \<in> Z \<Longrightarrow> Complete_Partial_Order.chain (\<le>) (f ` Y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   196
    by(rule chain_imageI[OF Y])(rule monotoneD[OF mono])
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   197
  have chain2: "Complete_Partial_Order.chain (\<le>) ((\<lambda>x. \<Squnion>(x ` Y)) ` Z)" using Z
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   198
  proof(rule chain_imageI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   199
    fix f g
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   200
    assume "f \<in> Z" "g \<in> Z"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   201
      and "fun_ord (\<le>) f g"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   202
    from chain1[OF \<open>f \<in> Z\<close>] show "\<Squnion>(f ` Y) \<le> \<Squnion>(g ` Y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   203
    proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   204
      fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   205
      assume "x \<in> f ` Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   206
      then obtain y where "y \<in> Y" "x = f y" by blast note this(2)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   207
      also have "\<dots> \<le> g y" using \<open>fun_ord (\<le>) f g\<close> by(simp add: fun_ord_def)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   208
      also have "\<dots> \<le> \<Squnion>(g ` Y)" using chain1[OF \<open>g \<in> Z\<close>]
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   209
        by(rule ccpo_Sup_upper)(simp add: \<open>y \<in> Y\<close>)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   210
      finally show "x \<le> \<Squnion>(g ` Y)" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   211
    qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   212
  qed
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   213
  have chain3: "\<And>x. Complete_Partial_Order.chain (\<le>) ((\<lambda>f. f x) ` Z)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   214
    using Z by(rule chain_imageI)(simp add: fun_ord_def)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   215
  have chain4: "Complete_Partial_Order.chain (\<le>) ((\<lambda>x. \<Squnion>((\<lambda>f. f x) ` Z)) ` Y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   216
    using Y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   217
  proof(rule chain_imageI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   218
    fix f x y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   219
    assume "x \<sqsubseteq> y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   220
    show "\<Squnion>((\<lambda>f. f x) ` Z) \<le> \<Squnion>((\<lambda>f. f y) ` Z)" (is "_ \<le> ?rhs") using chain3
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   221
    proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   222
      fix x'
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   223
      assume "x' \<in> (\<lambda>f. f x) ` Z"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   224
      then obtain f where "f \<in> Z" "x' = f x" by blast 
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   225
      then show "x' \<le> ?rhs"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   226
        by (metis (mono_tags, lifting) \<open>x \<sqsubseteq> y\<close> chain3 imageI ccpo_Sup_upper
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   227
            order_trans mono monotoneD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   228
    qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   229
  qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   230
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   231
  from chain2 have "?lhs \<le> ?rhs"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   232
  proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   233
    fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   234
    assume "x \<in> (\<lambda>x. \<Squnion>(x ` Y)) ` Z"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   235
    then obtain f where "f \<in> Z" "x = \<Squnion>(f ` Y)" by blast note this(2)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   236
    also have "\<dots> \<le> ?rhs" using chain1[OF \<open>f \<in> Z\<close>]
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   237
    proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   238
      fix x'
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   239
      assume "x' \<in> f ` Y"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   240
      then obtain y where "y \<in> Y" "x' = f y" by blast
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   241
      then show "x' \<le> ?rhs"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   242
        by (metis (mono_tags, lifting) \<open>f \<in> Z\<close> chain3 chain4 imageI local.ccpo_Sup_upper
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   243
            order.trans)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   244
    qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   245
    finally show "x \<le> ?rhs" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   246
  qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   247
  moreover
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   248
  have "?rhs \<le> ?lhs" using chain4
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   249
  proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   250
    fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   251
    assume "x \<in> (\<lambda>x. \<Squnion>((\<lambda>f. f x) ` Z)) ` Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   252
    then obtain y where "y \<in> Y" "x = \<Squnion>((\<lambda>f. f y) ` Z)" by blast note this(2)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   253
    also have "\<dots> \<le> ?lhs" using chain3
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   254
    proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   255
      fix x'
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   256
      assume "x' \<in> (\<lambda>f. f y) ` Z"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   257
      then obtain f where "f \<in> Z" "x' = f y" by blast
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   258
      then show "x' \<le> ?lhs"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   259
        by (metis (mono_tags, lifting) \<open>y \<in> Y\<close> ccpo_Sup_below_iff chain1 chain2 imageI
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   260
            ccpo_Sup_upper)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   261
    qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   262
    finally show "x \<le> ?lhs" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   263
  qed
73411
1f1366966296 avoid name clash
haftmann
parents: 70961
diff changeset
   264
  ultimately show "?lhs = ?rhs"
1f1366966296 avoid name clash
haftmann
parents: 70961
diff changeset
   265
    by (rule order.antisym)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   266
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   267
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   268
lemma fixp_mono:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   269
  assumes fg: "fun_ord (\<le>) f g"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   270
  and f: "monotone (\<le>) (\<le>) f"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   271
  and g: "monotone (\<le>) (\<le>) g"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   272
  shows "ccpo_class.fixp f \<le> ccpo_class.fixp g"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   273
unfolding fixp_def
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   274
proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   275
  fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   276
  assume "x \<in> ccpo_class.iterates f"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   277
  thus "x \<le> \<Squnion>ccpo_class.iterates g"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   278
  proof induction
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   279
    case (step x)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   280
    from f step.IH have "f x \<le> f (\<Squnion>ccpo_class.iterates g)" by(rule monotoneD)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   281
    also have "\<dots> \<le> g (\<Squnion>ccpo_class.iterates g)" using fg by(simp add: fun_ord_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   282
    also have "\<dots> = \<Squnion>ccpo_class.iterates g" by(fold fixp_def fixp_unfold[OF g]) simp
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   283
    finally show ?case .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   284
  qed(blast intro: ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   285
qed(rule chain_iterates[OF f])
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   286
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 76749
diff changeset
   287
context fixes ordb :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix \<open>\<sqsubseteq>\<close> 60) begin
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   288
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   289
lemma iterates_mono:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   290
  assumes f: "f \<in> ccpo.iterates (fun_lub Sup) (fun_ord (\<le>)) F"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   291
    and mono: "\<And>f. monotone (\<sqsubseteq>) (\<le>) f \<Longrightarrow> monotone (\<sqsubseteq>) (\<le>) (F f)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   292
  shows "monotone (\<sqsubseteq>) (\<le>) f"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   293
  using f
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   294
  by(induction rule: ccpo.iterates.induct[OF ccpo_fun, consumes 1])(blast intro: mono mono_lub)+
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   295
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   296
lemma fixp_preserves_mono:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   297
  assumes mono: "\<And>x. monotone (fun_ord (\<le>)) (\<le>) (\<lambda>f. F f x)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   298
  and mono2: "\<And>f. monotone (\<sqsubseteq>) (\<le>) f \<Longrightarrow> monotone (\<sqsubseteq>) (\<le>) (F f)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   299
  shows "monotone (\<sqsubseteq>) (\<le>) (ccpo.fixp (fun_lub Sup) (fun_ord (\<le>)) F)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   300
  (is "monotone _ _ ?fixp")
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   301
proof(rule monotoneI)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   302
  have mono: "monotone (fun_ord (\<le>)) (fun_ord (\<le>)) F"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   303
    by(rule monotoneI)(auto simp add: fun_ord_def intro: monotoneD[OF mono])
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   304
  let ?iter = "ccpo.iterates (fun_lub Sup) (fun_ord (\<le>)) F"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   305
  have chain: "\<And>x. Complete_Partial_Order.chain (\<le>) ((\<lambda>f. f x) ` ?iter)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   306
    by(rule chain_imageI[OF ccpo.chain_iterates[OF ccpo_fun mono]])(simp add: fun_ord_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   307
  fix x y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   308
  assume "x \<sqsubseteq> y"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   309
  have "(\<Squnion>f\<in>?iter. f x) \<le> (\<Squnion>f\<in>?iter. f y)"
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63092
diff changeset
   310
    using chain
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   311
  proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   312
    fix x'
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   313
    assume "x' \<in> (\<lambda>f. f x) ` ?iter"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   314
    then obtain f where f: "f \<in> ?iter" "x' = f x" by blast 
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   315
    then have "f x \<le> f y"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   316
      by (metis \<open>x \<sqsubseteq> y\<close> iterates_mono mono2 monotoneD)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   317
    also have "f y \<le> \<Squnion>((\<lambda>f. f y) ` ?iter)"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   318
      using chain f local.ccpo_Sup_upper by auto
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   319
    finally show "x' \<le> \<dots>"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   320
      using f(2) by blast 
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   321
  qed
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   322
  then show "?fixp x \<le> ?fixp y"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   323
    unfolding ccpo.fixp_def[OF ccpo_fun] fun_lub_apply .
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   324
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   325
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   326
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   327
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   328
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   329
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   330
lemma monotone2monotone:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   331
  assumes 2: "\<And>x. monotone ordb ordc (\<lambda>y. f x y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   332
  and t: "monotone orda ordb (\<lambda>x. t x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   333
  and 1: "\<And>y. monotone orda ordc (\<lambda>x. f x y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   334
  and trans: "transp ordc"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   335
shows "monotone orda ordc (\<lambda>x. f x (t x))"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   336
  using assms unfolding monotone_on_def by (metis UNIV_I transpE)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   337
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   338
subsection \<open>Continuity\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   339
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   340
definition cont :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b set \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   341
where
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   342
  "cont luba orda lubb ordb f \<longleftrightarrow> 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   343
  (\<forall>Y. Complete_Partial_Order.chain orda Y \<longrightarrow> Y \<noteq> {} \<longrightarrow> f (luba Y) = lubb (f ` Y))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   344
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   345
definition mcont :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b set \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   346
where
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   347
  "mcont luba orda lubb ordb f \<longleftrightarrow>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   348
   monotone orda ordb f \<and> cont luba orda lubb ordb f"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   349
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   350
subsubsection \<open>Theorem collection \<open>cont_intro\<close>\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   351
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   352
named_theorems cont_intro "continuity and admissibility intro rules"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   353
ML \<open>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   354
(* apply cont_intro rules as intro and try to solve 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   355
   the remaining of the emerging subgoals with simp *)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   356
fun cont_intro_tac ctxt =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69546
diff changeset
   357
  REPEAT_ALL_NEW (resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\<open>cont_intro\<close>)))
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   358
  THEN_ALL_NEW (SOLVED' (simp_tac ctxt))
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   359
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   360
fun cont_intro_simproc ctxt ct =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   361
  let
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   362
    fun mk_stmt t = t
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   363
      |> HOLogic.mk_Trueprop
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   364
      |> Thm.cterm_of ctxt
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   365
      |> Goal.init
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   366
    fun mk_thm t =
75650
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   367
      if exists_subterm Term.is_Var t then
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   368
        NONE
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   369
      else
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   370
        case SINGLE (cont_intro_tac ctxt 1) (mk_stmt t) of
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   371
          SOME thm => SOME (Goal.finish ctxt thm RS @{thm Eq_TrueI})
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   372
        | NONE => NONE
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   373
  in
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   374
    case Thm.term_of ct of
75582
6fb4a0829cc4 added predicate monotone_on and redefined monotone to be an abbreviation.
desharna
parents: 74334
diff changeset
   375
      t as \<^Const_>\<open>ccpo.admissible _ for _ _ _\<close> => mk_thm t
6fb4a0829cc4 added predicate monotone_on and redefined monotone to be an abbreviation.
desharna
parents: 74334
diff changeset
   376
    | t as \<^Const_>\<open>mcont _ _ for _ _ _ _ _\<close> => mk_thm t
75650
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   377
    | t as \<^Const_>\<open>monotone_on _ _ for _ _ _ _\<close> => mk_thm t
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   378
    | _ => NONE
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   379
  end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   380
  handle THM _ => NONE 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   381
  | TYPE _ => NONE
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   382
\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   383
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   384
simproc_setup "cont_intro"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   385
  ( "ccpo.admissible lub ord P"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   386
  | "mcont lub ord lub' ord' f"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   387
  | "monotone ord ord' f"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   388
  ) = \<open>K cont_intro_simproc\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   389
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   390
lemmas [cont_intro] =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   391
  call_mono
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   392
  let_mono
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   393
  if_mono
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   394
  option.const_mono
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   395
  tailrec.const_mono
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   396
  bind_mono
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   397
75650
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   398
experiment begin
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   399
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   400
text \<open>The following proof by simplification diverges if variables are not handled properly.\<close>
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   401
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   402
lemma "(\<And>f. monotone R S f \<Longrightarrow> thesis) \<Longrightarrow> monotone R S g \<Longrightarrow> thesis"
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   403
  by simp
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   404
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   405
end
6d4fb57eb66c fixed diverging simproc cont_intro
desharna
parents: 75582
diff changeset
   406
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   407
declare if_mono[simp]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   408
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   409
lemma monotone_id' [cont_intro]: "monotone ord ord (\<lambda>x. x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   410
  by(simp add: monotone_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   411
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   412
lemma monotone_applyI:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   413
  "monotone orda ordb F \<Longrightarrow> monotone (fun_ord orda) ordb (\<lambda>f. F (f x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   414
  by(rule monotoneI)(auto simp add: fun_ord_def dest: monotoneD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   415
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   416
lemma monotone_if_fun [partial_function_mono]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   417
  "\<lbrakk> monotone (fun_ord orda) (fun_ord ordb) F; monotone (fun_ord orda) (fun_ord ordb) G \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   418
  \<Longrightarrow> monotone (fun_ord orda) (fun_ord ordb) (\<lambda>f n. if c n then F f n else G f n)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   419
  by(simp add: monotone_def fun_ord_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   420
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   421
lemma monotone_fun_apply_fun [partial_function_mono]: 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   422
  "monotone (fun_ord (fun_ord ord)) (fun_ord ord) (\<lambda>f n. f t (g n))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   423
  by(rule monotoneI)(simp add: fun_ord_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   424
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   425
lemma monotone_fun_ord_apply: 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   426
  "monotone orda (fun_ord ordb) f \<longleftrightarrow> (\<forall>x. monotone orda ordb (\<lambda>y. f y x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   427
  by(auto simp add: monotone_def fun_ord_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   428
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   429
context preorder begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   430
76749
11a24dab1880 strengthened and renamed lemmas preorder.transp_(ge|gr|le|less)
desharna
parents: 75650
diff changeset
   431
declare transp_on_le[cont_intro]
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   432
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   433
lemma monotone_const [simp, cont_intro]: "monotone ord (\<le>) (\<lambda>_. c)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   434
  by(rule monotoneI) simp
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   435
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   436
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   437
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   438
lemma transp_le [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   439
  "class.preorder ord (mk_less ord) \<Longrightarrow> transp ord"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   440
  by(rule preorder.transp_on_le)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   441
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   442
context partial_function_definitions begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   443
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   444
declare const_mono [cont_intro, simp]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   445
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   446
lemma transp_le [cont_intro, simp]: "transp leq"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   447
  by(rule transpI)(rule leq_trans)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   448
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   449
lemma preorder [cont_intro, simp]: "class.preorder leq (mk_less leq)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   450
  by(unfold_locales)(auto simp add: mk_less_def intro: leq_refl leq_trans)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   451
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   452
declare ccpo[cont_intro, simp]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   453
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   454
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   455
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   456
lemma contI [intro?]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   457
  "(\<And>Y. \<lbrakk> Complete_Partial_Order.chain orda Y; Y \<noteq> {} \<rbrakk> \<Longrightarrow> f (luba Y) = lubb (f ` Y)) 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   458
  \<Longrightarrow> cont luba orda lubb ordb f"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   459
  unfolding cont_def by blast
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   460
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   461
lemma contD:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   462
  "\<lbrakk> cont luba orda lubb ordb f; Complete_Partial_Order.chain orda Y; Y \<noteq> {} \<rbrakk> 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   463
  \<Longrightarrow> f (luba Y) = lubb (f ` Y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   464
  unfolding cont_def by blast
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   465
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   466
lemma cont_id [simp, cont_intro]: "\<And>Sup. cont Sup ord Sup ord id"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   467
  by(rule contI) simp
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   468
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   469
lemma cont_id' [simp, cont_intro]: "\<And>Sup. cont Sup ord Sup ord (\<lambda>x. x)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   470
  by (simp add: Inf.INF_identity_eq contI)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   471
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   472
lemma cont_applyI [cont_intro]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   473
  assumes cont: "cont luba orda lubb ordb g"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   474
  shows "cont (fun_lub luba) (fun_ord orda) lubb ordb (\<lambda>f. g (f x))"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   475
  using assms by (simp add: cont_def chain_fun_ordD fun_lub_apply image_image)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   476
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   477
lemma call_cont: "cont (fun_lub lub) (fun_ord ord) lub ord (\<lambda>f. f t)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   478
  by(simp add: cont_def fun_lub_apply)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   479
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   480
lemma cont_if [cont_intro]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   481
  "\<lbrakk> cont luba orda lubb ordb f; cont luba orda lubb ordb g \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   482
  \<Longrightarrow> cont luba orda lubb ordb (\<lambda>x. if c then f x else g x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   483
  by(cases c) simp_all
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   484
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   485
lemma mcontI [intro?]:
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   486
  "\<lbrakk> monotone orda ordb f; cont luba orda lubb ordb f \<rbrakk> \<Longrightarrow> mcont luba orda lubb ordb f"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   487
  by(simp add: mcont_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   488
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   489
lemma mcont_mono: "mcont luba orda lubb ordb f \<Longrightarrow> monotone orda ordb f"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   490
  by(simp add: mcont_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   491
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   492
lemma mcont_cont [simp]: "mcont luba orda lubb ordb f \<Longrightarrow> cont luba orda lubb ordb f"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   493
  by(simp add: mcont_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   494
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   495
lemma mcont_monoD:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   496
  "\<lbrakk> mcont luba orda lubb ordb f; orda x y \<rbrakk> \<Longrightarrow> ordb (f x) (f y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   497
  by(auto simp add: mcont_def dest: monotoneD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   498
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   499
lemma mcont_contD:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   500
  "\<lbrakk> mcont luba orda lubb ordb f; Complete_Partial_Order.chain orda Y; Y \<noteq> {} \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   501
  \<Longrightarrow> f (luba Y) = lubb (f ` Y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   502
  by(auto simp add: mcont_def dest: contD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   503
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   504
lemma mcont_call [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   505
  "mcont (fun_lub lub) (fun_ord ord) lub ord (\<lambda>f. f t)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   506
  by(simp add: mcont_def call_mono call_cont)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   507
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   508
lemma mcont_id' [cont_intro, simp]: "mcont lub ord lub ord (\<lambda>x. x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   509
  by(simp add: mcont_def monotone_id')
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   510
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   511
lemma mcont_applyI:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   512
  "mcont luba orda lubb ordb (\<lambda>x. F x) \<Longrightarrow> mcont (fun_lub luba) (fun_ord orda) lubb ordb (\<lambda>f. F (f x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   513
  by(simp add: mcont_def monotone_applyI cont_applyI)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   514
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   515
lemma mcont_if [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   516
  "\<lbrakk> mcont luba orda lubb ordb (\<lambda>x. f x); mcont luba orda lubb ordb (\<lambda>x. g x) \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   517
  \<Longrightarrow> mcont luba orda lubb ordb (\<lambda>x. if c then f x else g x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   518
  by(simp add: mcont_def cont_if)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   519
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   520
lemma cont_fun_lub_apply: 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   521
  "cont luba orda (fun_lub lubb) (fun_ord ordb) f \<longleftrightarrow> (\<forall>x. cont luba orda lubb ordb (\<lambda>y. f y x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   522
  by(simp add: cont_def fun_lub_def fun_eq_iff)(auto simp add: image_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   523
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   524
lemma mcont_fun_lub_apply: 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   525
  "mcont luba orda (fun_lub lubb) (fun_ord ordb) f \<longleftrightarrow> (\<forall>x. mcont luba orda lubb ordb (\<lambda>y. f y x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   526
  by(auto simp add: monotone_fun_ord_apply cont_fun_lub_apply mcont_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   527
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   528
context ccpo begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   529
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   530
lemma cont_const [simp, cont_intro]: "cont luba orda Sup (\<le>) (\<lambda>x. c)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   531
  by (rule contI) (simp add: image_constant_conv cong del: SUP_cong_simp)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   532
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   533
lemma mcont_const [cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   534
  "mcont luba orda Sup (\<le>) (\<lambda>x. c)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   535
  by(simp add: mcont_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   536
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   537
lemma cont_apply:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   538
  assumes 2: "\<And>x. cont lubb ordb Sup (\<le>) (\<lambda>y. f x y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   539
    and t: "cont luba orda lubb ordb (\<lambda>x. t x)"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   540
    and 1: "\<And>y. cont luba orda Sup (\<le>) (\<lambda>x. f x y)"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   541
    and mono: "monotone orda ordb (\<lambda>x. t x)"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   542
    and mono2: "\<And>x. monotone ordb (\<le>) (\<lambda>y. f x y)"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   543
    and mono1: "\<And>y. monotone orda (\<le>) (\<lambda>x. f x y)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   544
  shows "cont luba orda Sup (\<le>) (\<lambda>x. f x (t x))"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   545
proof
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   546
  fix Y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   547
  assume chain: "Complete_Partial_Order.chain orda Y" and "Y \<noteq> {}"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   548
  moreover from chain have chain': "Complete_Partial_Order.chain ordb (t ` Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   549
    by(rule chain_imageI)(rule monotoneD[OF mono])
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   550
  ultimately show "f (luba Y) (t (luba Y)) = \<Squnion>((\<lambda>x. f x (t x)) ` Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   551
    by(simp add: contD[OF 1] contD[OF t] contD[OF 2] image_image)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   552
      (rule diag_Sup[OF chain], auto intro: monotone2monotone[OF mono2 mono monotone_const transpI] monotoneD[OF mono1])
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   553
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   554
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   555
lemma mcont2mcont':
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   556
  "\<lbrakk> \<And>x. mcont lub' ord' Sup (\<le>) (\<lambda>y. f x y);
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   557
     \<And>y. mcont lub ord Sup (\<le>) (\<lambda>x. f x y);
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   558
     mcont lub ord lub' ord' (\<lambda>y. t y) \<rbrakk>
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   559
  \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. f x (t x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   560
  unfolding mcont_def by(blast intro: transp_on_le monotone2monotone cont_apply)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   561
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   562
lemma mcont2mcont:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   563
  "\<lbrakk>mcont lub' ord' Sup (\<le>) (\<lambda>x. f x); mcont lub ord lub' ord' (\<lambda>x. t x)\<rbrakk> 
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   564
  \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. f (t x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   565
  by(rule mcont2mcont'[OF _ mcont_const]) 
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   566
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   567
context
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 76749
diff changeset
   568
  fixes ord :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix \<open>\<sqsubseteq>\<close> 60) 
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   569
    and lub :: "'b set \<Rightarrow> 'b" (\<open>\<Or>\<close>)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   570
begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   571
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   572
lemma cont_fun_lub_Sup:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   573
  assumes chainM: "Complete_Partial_Order.chain (fun_ord (\<le>)) M"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   574
  and mcont [rule_format]: "\<forall>f\<in>M. mcont lub (\<sqsubseteq>) Sup (\<le>) f"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   575
  shows "cont lub (\<sqsubseteq>) Sup (\<le>) (fun_lub Sup M)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   576
proof(rule contI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   577
  fix Y
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   578
  assume chain: "Complete_Partial_Order.chain (\<sqsubseteq>) Y"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   579
    and Y: "Y \<noteq> {}"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   580
  from swap_Sup[OF chain chainM mcont[THEN mcont_mono]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   581
  show "fun_lub Sup M (\<Or>Y) = \<Squnion>(fun_lub Sup M ` Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   582
    by(simp add: mcont_contD[OF mcont chain Y] fun_lub_apply cong: image_cong)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   583
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   584
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   585
lemma mcont_fun_lub_Sup:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   586
  "\<lbrakk> Complete_Partial_Order.chain (fun_ord (\<le>)) M;
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   587
    \<forall>f\<in>M. mcont lub ord Sup (\<le>) f \<rbrakk>
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   588
  \<Longrightarrow> mcont lub (\<sqsubseteq>) Sup (\<le>) (fun_lub Sup M)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   589
by(simp add: mcont_def cont_fun_lub_Sup mono_lub)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   590
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   591
lemma iterates_mcont:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   592
  assumes f: "f \<in> ccpo.iterates (fun_lub Sup) (fun_ord (\<le>)) F"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   593
  and mono: "\<And>f. mcont lub (\<sqsubseteq>) Sup (\<le>) f \<Longrightarrow> mcont lub (\<sqsubseteq>) Sup (\<le>) (F f)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   594
  shows "mcont lub (\<sqsubseteq>) Sup (\<le>) f"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   595
using f
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   596
by(induction rule: ccpo.iterates.induct[OF ccpo_fun, consumes 1, case_names step Sup])(blast intro: mono mcont_fun_lub_Sup)+
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   597
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   598
lemma fixp_preserves_mcont:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   599
  assumes mono: "\<And>x. monotone (fun_ord (\<le>)) (\<le>) (\<lambda>f. F f x)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   600
  and mcont: "\<And>f. mcont lub (\<sqsubseteq>) Sup (\<le>) f \<Longrightarrow> mcont lub (\<sqsubseteq>) Sup (\<le>) (F f)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   601
  shows "mcont lub (\<sqsubseteq>) Sup (\<le>) (ccpo.fixp (fun_lub Sup) (fun_ord (\<le>)) F)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   602
  (is "mcont _ _ _ _ ?fixp")
81281
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   603
  unfolding mcont_def
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   604
proof(intro conjI monotoneI contI)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   605
  have mono: "monotone (fun_ord (\<le>)) (fun_ord (\<le>)) F"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   606
    by(rule monotoneI)(auto simp add: fun_ord_def intro: monotoneD[OF mono])
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   607
  let ?iter = "ccpo.iterates (fun_lub Sup) (fun_ord (\<le>)) F"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   608
  have chain: "\<And>x. Complete_Partial_Order.chain (\<le>) ((\<lambda>f. f x) ` ?iter)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   609
    by(rule chain_imageI[OF ccpo.chain_iterates[OF ccpo_fun mono]])(simp add: fun_ord_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   610
81281
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   611
  show "?fixp x \<le> ?fixp y" if "x \<sqsubseteq> y" for x y
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   612
  proof -
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   613
    have "(\<Squnion>f\<in>?iter. f x)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   614
    \<le> (\<Squnion>f\<in>?iter. f y)"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   615
      using chain
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   616
    proof(rule ccpo_Sup_least)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   617
      fix x'
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   618
      assume "x' \<in> (\<lambda>f. f x) ` ?iter"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   619
      then obtain f where f: "f \<in> ?iter" "x' = f x" by blast 
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   620
      then have "f x \<le> f y"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   621
        by (metis iterates_mcont mcont mcont_monoD that)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   622
      also have "f y \<le> \<Squnion>((\<lambda>f. f y) ` ?iter)"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   623
        using chain f local.ccpo_Sup_upper by auto
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   624
      finally show "x' \<le> \<dots>"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   625
        using f(2) by blast 
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   626
    qed
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   627
    then show ?thesis
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   628
      by (simp add: ccpo.fixp_def[OF ccpo_fun] fun_lub_apply)
81281
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   629
  qed
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   630
  show "?fixp (\<Or>Y) = \<Squnion>(?fixp ` Y)"
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   631
    if chain: "Complete_Partial_Order.chain (\<sqsubseteq>) Y" and Y: "Y \<noteq> {}" for Y
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   632
  proof -
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   633
    have "f (\<Or>Y) = \<Squnion>(f ` Y)" if "f \<in> ?iter" for f
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   634
      using that mcont chain Y
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   635
      by (rule mcont_contD[OF iterates_mcont])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   636
    moreover have "\<Squnion>((\<lambda>f. \<Squnion>(f ` Y)) ` ?iter) = \<Squnion>((\<lambda>x. \<Squnion>((\<lambda>f. f x) ` ?iter)) ` Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   637
      using chain ccpo.chain_iterates[OF ccpo_fun mono]
81281
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   638
      by (rule swap_Sup)(rule mcont_mono[OF iterates_mcont[OF _ mcont]])
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   639
    ultimately show ?thesis
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   640
      unfolding ccpo.fixp_def[OF ccpo_fun]
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   641
      by (simp add: fun_lub_apply cong: image_cong)
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
   642
  qed
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   643
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   644
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   645
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   646
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   647
context
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   648
  fixes F :: "'c \<Rightarrow> 'c" and U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a" and C :: "('b \<Rightarrow> 'a) \<Rightarrow> 'c" and f
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   649
  assumes mono: "\<And>x. monotone (fun_ord (\<le>)) (\<le>) (\<lambda>f. U (F (C f)) x)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   650
  and eq: "f \<equiv> C (ccpo.fixp (fun_lub Sup) (fun_ord (\<le>)) (\<lambda>f. U (F (C f))))"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   651
  and inverse: "\<And>f. U (C f) = f"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   652
begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   653
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   654
lemma fixp_preserves_mono_uc:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   655
  assumes mono2: "\<And>f. monotone ord (\<le>) (U f) \<Longrightarrow> monotone ord (\<le>) (U (F f))"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   656
  shows "monotone ord (\<le>) (U f)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   657
using fixp_preserves_mono[OF mono mono2] by(subst eq)(simp add: inverse)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   658
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   659
lemma fixp_preserves_mcont_uc:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   660
  assumes mcont: "\<And>f. mcont lubb ordb Sup (\<le>) (U f) \<Longrightarrow> mcont lubb ordb Sup (\<le>) (U (F f))"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   661
  shows "mcont lubb ordb Sup (\<le>) (U f)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   662
using fixp_preserves_mcont[OF mono mcont] by(subst eq)(simp add: inverse)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   663
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   664
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   665
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   666
lemmas fixp_preserves_mono1 = fixp_preserves_mono_uc[of "\<lambda>x. x" _ "\<lambda>x. x", OF _ _ refl]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   667
lemmas fixp_preserves_mono2 =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   668
  fixp_preserves_mono_uc[of "case_prod" _ "curry", unfolded case_prod_curry curry_case_prod, OF _ _ refl]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   669
lemmas fixp_preserves_mono3 =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   670
  fixp_preserves_mono_uc[of "\<lambda>f. case_prod (case_prod f)" _ "\<lambda>f. curry (curry f)", unfolded case_prod_curry curry_case_prod, OF _ _ refl]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   671
lemmas fixp_preserves_mono4 =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   672
  fixp_preserves_mono_uc[of "\<lambda>f. case_prod (case_prod (case_prod f))" _ "\<lambda>f. curry (curry (curry f))", unfolded case_prod_curry curry_case_prod, OF _ _ refl]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   673
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   674
lemmas fixp_preserves_mcont1 = fixp_preserves_mcont_uc[of "\<lambda>x. x" _ "\<lambda>x. x", OF _ _ refl]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   675
lemmas fixp_preserves_mcont2 =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   676
  fixp_preserves_mcont_uc[of "case_prod" _ "curry", unfolded case_prod_curry curry_case_prod, OF _ _ refl]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   677
lemmas fixp_preserves_mcont3 =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   678
  fixp_preserves_mcont_uc[of "\<lambda>f. case_prod (case_prod f)" _ "\<lambda>f. curry (curry f)", unfolded case_prod_curry curry_case_prod, OF _ _ refl]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   679
lemmas fixp_preserves_mcont4 =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   680
  fixp_preserves_mcont_uc[of "\<lambda>f. case_prod (case_prod (case_prod f))" _ "\<lambda>f. curry (curry (curry f))", unfolded case_prod_curry curry_case_prod, OF _ _ refl]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   681
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   682
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   683
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   684
lemma (in preorder) monotone_if_bot:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   685
  fixes bot
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   686
  assumes mono: "\<And>x y. \<lbrakk> x \<le> y; \<not> (x \<le> bound) \<rbrakk> \<Longrightarrow> ord (f x) (f y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   687
  and bot: "\<And>x. \<not> x \<le> bound \<Longrightarrow> ord bot (f x)" "ord bot bot"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   688
  shows "monotone (\<le>) ord (\<lambda>x. if x \<le> bound then bot else f x)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   689
by(rule monotoneI)(auto intro: bot intro: mono order_trans)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   690
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   691
lemma (in ccpo) mcont_if_bot:
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 76749
diff changeset
   692
  fixes bot and lub (\<open>\<Or>\<close>) and ord (infix \<open>\<sqsubseteq>\<close> 60)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   693
  assumes ccpo: "class.ccpo lub (\<sqsubseteq>) lt"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   694
  and mono: "\<And>x y. \<lbrakk> x \<le> y; \<not> x \<le> bound \<rbrakk> \<Longrightarrow> f x \<sqsubseteq> f y"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   695
  and cont: "\<And>Y. \<lbrakk> Complete_Partial_Order.chain (\<le>) Y; Y \<noteq> {}; \<And>x. x \<in> Y \<Longrightarrow> \<not> x \<le> bound \<rbrakk> \<Longrightarrow> f (\<Squnion>Y) = \<Or>(f ` Y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   696
  and bot: "\<And>x. \<not> x \<le> bound \<Longrightarrow> bot \<sqsubseteq> f x"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   697
  shows "mcont Sup (\<le>) lub (\<sqsubseteq>) (\<lambda>x. if x \<le> bound then bot else f x)" (is "mcont _ _ _ _ ?g")
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   698
proof(intro mcontI contI)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   699
  interpret c: ccpo lub "(\<sqsubseteq>)" lt by(fact ccpo)
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   700
  show "monotone (\<le>) (\<sqsubseteq>) ?g" by(rule monotone_if_bot)(simp_all add: mono bot)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   701
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   702
  fix Y
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   703
  assume chain: "Complete_Partial_Order.chain (\<le>) Y" and Y: "Y \<noteq> {}"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   704
  show "?g (\<Squnion>Y) = \<Or>(?g ` Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   705
  proof(cases "Y \<subseteq> {x. x \<le> bound}")
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   706
    case True
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   707
    hence "\<Squnion>Y \<le> bound" using chain by(auto intro: ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   708
    moreover have "Y \<inter> {x. \<not> x \<le> bound} = {}" using True by auto
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   709
    ultimately show ?thesis using True Y
69546
27dae626822b prefer naming convention from datatype package for strong congruence rules
haftmann
parents: 69164
diff changeset
   710
      by (auto simp add: image_constant_conv cong del: c.SUP_cong_simp)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   711
  next
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   712
    case False
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   713
    let ?Y = "Y \<inter> {x. \<not> x \<le> bound}"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   714
    have chain': "Complete_Partial_Order.chain (\<le>) ?Y"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   715
      using chain by(rule chain_subset) simp
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   716
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   717
    from False obtain y where ybound: "\<not> y \<le> bound" and y: "y \<in> Y" by blast
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   718
    hence "\<not> \<Squnion>Y \<le> bound" by (metis ccpo_Sup_upper chain order.trans)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   719
    hence "?g (\<Squnion>Y) = f (\<Squnion>Y)" by simp
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   720
    also have "\<Squnion>Y \<le> \<Squnion>?Y" using chain
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   721
    proof(rule ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   722
      fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   723
      assume x: "x \<in> Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   724
      show "x \<le> \<Squnion>?Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   725
      proof(cases "x \<le> bound")
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   726
        case True
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   727
        with chainD[OF chain x y] have "x \<le> y" using ybound by(auto intro: order_trans)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   728
        thus ?thesis by(rule order_trans)(auto intro: ccpo_Sup_upper[OF chain'] simp add: y ybound)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   729
      qed(auto intro: ccpo_Sup_upper[OF chain'] simp add: x)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   730
    qed
73411
1f1366966296 avoid name clash
haftmann
parents: 70961
diff changeset
   731
    hence "\<Squnion>Y = \<Squnion>?Y" by(rule order.antisym)(blast intro: ccpo_Sup_least[OF chain'] ccpo_Sup_upper[OF chain])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   732
    hence "f (\<Squnion>Y) = f (\<Squnion>?Y)" by simp
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   733
    also have "f (\<Squnion>?Y) = \<Or>(f ` ?Y)" using chain' by(rule cont)(insert y ybound, auto)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   734
    also have "\<Or>(f ` ?Y) = \<Or>(?g ` Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   735
    proof(cases "Y \<inter> {x. x \<le> bound} = {}")
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   736
      case True
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   737
      hence "f ` ?Y = ?g ` Y" by auto
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   738
      thus ?thesis by(rule arg_cong)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   739
    next
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   740
      case False
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   741
      have chain'': "Complete_Partial_Order.chain (\<sqsubseteq>) (insert bot (f ` ?Y))"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   742
        using chain by(auto intro!: chainI bot dest: chainD intro: mono)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   743
      hence chain''': "Complete_Partial_Order.chain (\<sqsubseteq>) (f ` ?Y)" by(rule chain_subset) blast
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   744
      have "bot \<sqsubseteq> \<Or>(f ` ?Y)" using y ybound by(blast intro: c.order_trans[OF bot] c.ccpo_Sup_upper[OF chain'''])
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   745
      hence "\<Or>(insert bot (f ` ?Y)) \<sqsubseteq> \<Or>(f ` ?Y)" using chain''
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   746
        by(auto intro: c.ccpo_Sup_least c.ccpo_Sup_upper[OF chain''']) 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   747
      with _ have "\<dots> = \<Or>(insert bot (f ` ?Y))"
73411
1f1366966296 avoid name clash
haftmann
parents: 70961
diff changeset
   748
        by(rule c.order.antisym)(blast intro: c.ccpo_Sup_least[OF chain'''] c.ccpo_Sup_upper[OF chain''])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   749
      also have "insert bot (f ` ?Y) = ?g ` Y" using False by auto
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   750
      finally show ?thesis .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   751
    qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   752
    finally show ?thesis .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   753
  qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   754
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   755
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   756
context partial_function_definitions begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   757
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   758
lemma mcont_const [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   759
  "mcont luba orda lub leq (\<lambda>x. c)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   760
by(rule ccpo.mcont_const)(rule Partial_Function.ccpo[OF partial_function_definitions_axioms])
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   761
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   762
lemmas [cont_intro, simp] =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   763
  ccpo.cont_const[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   764
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   765
lemma mono2mono:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   766
  assumes "monotone ordb leq (\<lambda>y. f y)" "monotone orda ordb (\<lambda>x. t x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   767
  shows "monotone orda leq (\<lambda>x. f (t x))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   768
using assms by(rule monotone2monotone) simp_all
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   769
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   770
lemmas mcont2mcont' = ccpo.mcont2mcont'[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   771
lemmas mcont2mcont = ccpo.mcont2mcont[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   772
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   773
lemmas fixp_preserves_mono1 = ccpo.fixp_preserves_mono1[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   774
lemmas fixp_preserves_mono2 = ccpo.fixp_preserves_mono2[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   775
lemmas fixp_preserves_mono3 = ccpo.fixp_preserves_mono3[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   776
lemmas fixp_preserves_mono4 = ccpo.fixp_preserves_mono4[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   777
lemmas fixp_preserves_mcont1 = ccpo.fixp_preserves_mcont1[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   778
lemmas fixp_preserves_mcont2 = ccpo.fixp_preserves_mcont2[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   779
lemmas fixp_preserves_mcont3 = ccpo.fixp_preserves_mcont3[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   780
lemmas fixp_preserves_mcont4 = ccpo.fixp_preserves_mcont4[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   781
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   782
lemma monotone_if_bot:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   783
  fixes bot
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   784
  assumes g: "\<And>x. g x = (if leq x bound then bot else f x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   785
  and mono: "\<And>x y. \<lbrakk> leq x y; \<not> leq x bound \<rbrakk> \<Longrightarrow> ord (f x) (f y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   786
  and bot: "\<And>x. \<not> leq x bound \<Longrightarrow> ord bot (f x)" "ord bot bot"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   787
  shows "monotone leq ord g"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   788
unfolding g[abs_def] using preorder mono bot by(rule preorder.monotone_if_bot)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   789
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   790
lemma mcont_if_bot:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   791
  fixes bot
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   792
  assumes ccpo: "class.ccpo lub' ord (mk_less ord)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   793
  and bot: "\<And>x. \<not> leq x bound \<Longrightarrow> ord bot (f x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   794
  and g: "\<And>x. g x = (if leq x bound then bot else f x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   795
  and mono: "\<And>x y. \<lbrakk> leq x y; \<not> leq x bound \<rbrakk> \<Longrightarrow> ord (f x) (f y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   796
  and cont: "\<And>Y. \<lbrakk> Complete_Partial_Order.chain leq Y; Y \<noteq> {}; \<And>x. x \<in> Y \<Longrightarrow> \<not> leq x bound \<rbrakk> \<Longrightarrow> f (lub Y) = lub' (f ` Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   797
  shows "mcont lub leq lub' ord g"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   798
unfolding g[abs_def] using ccpo mono cont bot by(rule ccpo.mcont_if_bot[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]])
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   799
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   800
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   801
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   802
subsection \<open>Admissibility\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   803
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   804
lemma admissible_subst:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   805
  assumes adm: "ccpo.admissible luba orda (\<lambda>x. P x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   806
  and mcont: "mcont lubb ordb luba orda f"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   807
shows "ccpo.admissible lubb ordb (\<lambda>x. P (f x))"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   808
  using assms by (simp add: ccpo.admissible_def chain_imageI mcont_contD mcont_monoD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   809
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   810
lemmas [simp, cont_intro] = 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   811
  admissible_all
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   812
  admissible_ball
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   813
  admissible_const
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   814
  admissible_conj
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   815
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   816
lemma admissible_disj' [simp, cont_intro]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   817
  "\<lbrakk> class.ccpo lub ord (mk_less ord); ccpo.admissible lub ord P; ccpo.admissible lub ord Q \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   818
  \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. P x \<or> Q x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   819
  by(rule ccpo.admissible_disj)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   820
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   821
lemma admissible_imp' [cont_intro]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   822
  "\<lbrakk> class.ccpo lub ord (mk_less ord);
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   823
     ccpo.admissible lub ord (\<lambda>x. \<not> P x);
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   824
     ccpo.admissible lub ord (\<lambda>x. Q x) \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   825
  \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. P x \<longrightarrow> Q x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   826
  unfolding imp_conv_disj by(rule ccpo.admissible_disj)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   827
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   828
lemma admissible_imp [cont_intro]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   829
  "(Q \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. P x))
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   830
  \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. Q \<longrightarrow> P x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   831
  by(rule ccpo.admissibleI)(auto dest: ccpo.admissibleD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   832
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   833
lemma admissible_not_mem' [THEN admissible_subst, cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   834
  shows admissible_not_mem: "ccpo.admissible Union (\<subseteq>) (\<lambda>A. x \<notin> A)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   835
  by(rule ccpo.admissibleI) auto
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   836
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   837
lemma admissible_eqI:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   838
  assumes f: "cont luba orda lub ord (\<lambda>x. f x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   839
    and g: "cont luba orda lub ord (\<lambda>x. g x)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   840
  shows "ccpo.admissible luba orda (\<lambda>x. f x = g x)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
   841
  by (smt (verit, best) Sup.SUP_cong ccpo.admissible_def contD assms)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   842
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   843
corollary admissible_eq_mcontI [cont_intro]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   844
  "\<lbrakk> mcont luba orda lub ord (\<lambda>x. f x); 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   845
    mcont luba orda lub ord (\<lambda>x. g x) \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   846
  \<Longrightarrow> ccpo.admissible luba orda (\<lambda>x. f x = g x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   847
  by(rule admissible_eqI)(auto simp add: mcont_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   848
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   849
lemma admissible_iff [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   850
  "\<lbrakk> ccpo.admissible lub ord (\<lambda>x. P x \<longrightarrow> Q x); ccpo.admissible lub ord (\<lambda>x. Q x \<longrightarrow> P x) \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   851
  \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. P x \<longleftrightarrow> Q x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   852
  by(subst iff_conv_conj_imp)(rule admissible_conj)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   853
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   854
context ccpo begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   855
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   856
lemma admissible_leI:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   857
  assumes f: "mcont luba orda Sup (\<le>) (\<lambda>x. f x)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   858
  and g: "mcont luba orda Sup (\<le>) (\<lambda>x. g x)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   859
  shows "ccpo.admissible luba orda (\<lambda>x. f x \<le> g x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   860
proof(rule ccpo.admissibleI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   861
  fix A
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   862
  assume chain: "Complete_Partial_Order.chain orda A"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   863
    and le: "\<forall>x\<in>A. f x \<le> g x"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   864
    and False: "A \<noteq> {}"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   865
  have "f (luba A) = \<Squnion>(f ` A)" by(simp add: mcont_contD[OF f] chain False)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   866
  also have "\<dots> \<le> \<Squnion>(g ` A)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   867
  proof(rule ccpo_Sup_least)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   868
    from chain show "Complete_Partial_Order.chain (\<le>) (f ` A)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   869
      by(rule chain_imageI)(rule mcont_monoD[OF f])
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   870
    fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   871
    assume "x \<in> f ` A"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   872
    then obtain y where "y \<in> A" "x = f y" by blast note this(2)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   873
    also have "f y \<le> g y" using le \<open>y \<in> A\<close> by simp
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   874
    also have "Complete_Partial_Order.chain (\<le>) (g ` A)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   875
      using chain by(rule chain_imageI)(rule mcont_monoD[OF g])
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
   876
    hence "g y \<le> \<Squnion>(g ` A)" by(rule ccpo_Sup_upper)(simp add: \<open>y \<in> A\<close>)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   877
    finally show "x \<le> \<dots>" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   878
  qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   879
  also have "\<dots> = g (luba A)" by(simp add: mcont_contD[OF g] chain False)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   880
  finally show "f (luba A) \<le> g (luba A)" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   881
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   882
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   883
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   884
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   885
lemma admissible_leI:
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 76749
diff changeset
   886
  fixes ord (infix \<open>\<sqsubseteq>\<close> 60) and lub (\<open>\<Or>\<close>)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   887
  assumes "class.ccpo lub (\<sqsubseteq>) (mk_less (\<sqsubseteq>))"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   888
  and "mcont luba orda lub (\<sqsubseteq>) (\<lambda>x. f x)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   889
  and "mcont luba orda lub (\<sqsubseteq>) (\<lambda>x. g x)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   890
  shows "ccpo.admissible luba orda (\<lambda>x. f x \<sqsubseteq> g x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   891
using assms by(rule ccpo.admissible_leI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   892
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   893
declare ccpo_class.admissible_leI[cont_intro]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   894
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   895
context ccpo begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   896
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   897
lemma admissible_not_below: "ccpo.admissible Sup (\<le>) (\<lambda>x. \<not> (\<le>) x y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   898
by(rule ccpo.admissibleI)(simp add: ccpo_Sup_below_iff)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   899
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   900
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   901
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   902
lemma (in preorder) preorder [cont_intro, simp]: "class.preorder (\<le>) (mk_less (\<le>))"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   903
by(unfold_locales)(auto simp add: mk_less_def intro: order_trans)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   904
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   905
context partial_function_definitions begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   906
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   907
lemmas [cont_intro, simp] =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   908
  admissible_leI[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   909
  ccpo.admissible_not_below[THEN admissible_subst, OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   910
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   911
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   912
66244
4c999b5d78e2 qualify Complete_Partial_Order2.compact
Andreas Lochbihler
parents: 65366
diff changeset
   913
setup \<open>Sign.map_naming (Name_Space.mandatory_path "ccpo")\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   914
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   915
inductive compact :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   916
  for lub ord x 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   917
where compact:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   918
  "\<lbrakk> ccpo.admissible lub ord (\<lambda>y. \<not> ord x y);
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   919
     ccpo.admissible lub ord (\<lambda>y. x \<noteq> y) \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   920
  \<Longrightarrow> compact lub ord x"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   921
66244
4c999b5d78e2 qualify Complete_Partial_Order2.compact
Andreas Lochbihler
parents: 65366
diff changeset
   922
setup \<open>Sign.map_naming Name_Space.parent_path\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   923
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   924
context ccpo begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   925
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   926
lemma compactI:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   927
  assumes "ccpo.admissible Sup (\<le>) (\<lambda>y. \<not> x \<le> y)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   928
  shows "ccpo.compact Sup (\<le>) x"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   929
using assms
66244
4c999b5d78e2 qualify Complete_Partial_Order2.compact
Andreas Lochbihler
parents: 65366
diff changeset
   930
proof(rule ccpo.compact.intros)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   931
  have neq: "(\<lambda>y. x \<noteq> y) = (\<lambda>y. \<not> x \<le> y \<or> \<not> y \<le> x)" by(auto)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   932
  show "ccpo.admissible Sup (\<le>) (\<lambda>y. x \<noteq> y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   933
    by(subst neq)(rule admissible_disj admissible_not_below assms)+
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   934
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   935
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   936
lemma compact_bot:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   937
  assumes "x = Sup {}"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   938
  shows "ccpo.compact Sup (\<le>) x"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   939
proof(rule compactI)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   940
  show "ccpo.admissible Sup (\<le>) (\<lambda>y. \<not> x \<le> y)" using assms
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   941
    by(auto intro!: ccpo.admissibleI intro: ccpo_Sup_least chain_empty)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   942
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   943
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   944
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   945
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   946
lemma admissible_compact_neq' [THEN admissible_subst, cont_intro, simp]:
66244
4c999b5d78e2 qualify Complete_Partial_Order2.compact
Andreas Lochbihler
parents: 65366
diff changeset
   947
  shows admissible_compact_neq: "ccpo.compact lub ord k \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. k \<noteq> x)"
4c999b5d78e2 qualify Complete_Partial_Order2.compact
Andreas Lochbihler
parents: 65366
diff changeset
   948
by(simp add: ccpo.compact.simps)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   949
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   950
lemma admissible_neq_compact' [THEN admissible_subst, cont_intro, simp]:
66244
4c999b5d78e2 qualify Complete_Partial_Order2.compact
Andreas Lochbihler
parents: 65366
diff changeset
   951
  shows admissible_neq_compact: "ccpo.compact lub ord k \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. x \<noteq> k)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   952
by(subst eq_commute)(rule admissible_compact_neq)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   953
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   954
context partial_function_definitions begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   955
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   956
lemmas [cont_intro, simp] = ccpo.compact_bot[OF Partial_Function.ccpo[OF partial_function_definitions_axioms]]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   957
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   958
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   959
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   960
context ccpo begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   961
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   962
lemma fixp_strong_induct:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   963
  assumes [cont_intro]: "ccpo.admissible Sup (\<le>) P"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   964
  and mono: "monotone (\<le>) (\<le>) f"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   965
  and bot: "P (\<Squnion>{})"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   966
  and step: "\<And>x. \<lbrakk> x \<le> ccpo_class.fixp f; P x \<rbrakk> \<Longrightarrow> P (f x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   967
  shows "P (ccpo_class.fixp f)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   968
proof(rule fixp_induct[where P="\<lambda>x. x \<le> ccpo_class.fixp f \<and> P x", THEN conjunct2])
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   969
  note [cont_intro] = admissible_leI
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
   970
  show "ccpo.admissible Sup (\<le>) (\<lambda>x. x \<le> ccpo_class.fixp f \<and> P x)" by simp
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   971
next
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   972
  show "\<Squnion>{} \<le> ccpo_class.fixp f \<and> P (\<Squnion>{})"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   973
    by(auto simp add: bot intro: ccpo_Sup_least chain_empty)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   974
next
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   975
  fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   976
  assume "x \<le> ccpo_class.fixp f \<and> P x"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   977
  thus "f x \<le> ccpo_class.fixp f \<and> P (f x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   978
    by(subst fixp_unfold[OF mono])(auto dest: monotoneD[OF mono] intro: step)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   979
qed(rule mono)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   980
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   981
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   982
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   983
context partial_function_definitions begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   984
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   985
lemma fixp_strong_induct_uc:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   986
  fixes F :: "'c \<Rightarrow> 'c"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   987
    and U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   988
    and C :: "('b \<Rightarrow> 'a) \<Rightarrow> 'c"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   989
    and P :: "('b \<Rightarrow> 'a) \<Rightarrow> bool"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   990
  assumes mono: "\<And>x. mono_body (\<lambda>f. U (F (C f)) x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   991
    and eq: "f \<equiv> C (fixp_fun (\<lambda>f. U (F (C f))))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   992
    and inverse: "\<And>f. U (C f) = f"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   993
    and adm: "ccpo.admissible lub_fun le_fun P"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   994
    and bot: "P (\<lambda>_. lub {})"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   995
    and step: "\<And>f'. \<lbrakk> P (U f'); le_fun (U f') (U f) \<rbrakk> \<Longrightarrow> P (U (F f'))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
   996
  shows "P (U f)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   997
  unfolding eq inverse
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   998
  apply (rule ccpo.fixp_strong_induct[OF ccpo adm])
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
   999
    apply (insert mono, auto simp: monotone_def fun_ord_def bot fun_lub_def)[2]
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1000
  apply (rule_tac f'5="C x" in step)
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1001
   apply (simp_all add: inverse eq)
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1002
  done
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1003
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1004
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1005
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69546
diff changeset
  1006
subsection \<open>\<^term>\<open>(=)\<close> as order\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1007
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1008
definition lub_singleton :: "('a set \<Rightarrow> 'a) \<Rightarrow> bool"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1009
  where "lub_singleton lub \<longleftrightarrow> (\<forall>a. lub {a} = a)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1010
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1011
definition the_Sup :: "'a set \<Rightarrow> 'a"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1012
  where "the_Sup A = (THE a. a \<in> A)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1013
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1014
lemma lub_singleton_the_Sup [cont_intro, simp]: "lub_singleton the_Sup"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1015
  by(simp add: lub_singleton_def the_Sup_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1016
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1017
lemma (in ccpo) lub_singleton: "lub_singleton Sup"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1018
  by(simp add: lub_singleton_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1019
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1020
lemma (in partial_function_definitions) lub_singleton [cont_intro, simp]: "lub_singleton lub"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1021
  by(rule ccpo.lub_singleton)(rule Partial_Function.ccpo[OF partial_function_definitions_axioms])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1022
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1023
lemma preorder_eq [cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1024
  "class.preorder (=) (mk_less (=))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1025
  by(unfold_locales)(simp_all add: mk_less_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1026
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1027
lemma monotone_eqI [cont_intro]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1028
  assumes "class.preorder ord (mk_less ord)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1029
  shows "monotone (=) ord f"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1030
proof -
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1031
  interpret preorder ord "mk_less ord" by fact
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1032
  show ?thesis by(simp add: monotone_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1033
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1034
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1035
lemma cont_eqI [cont_intro]: 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1036
  fixes f :: "'a \<Rightarrow> 'b"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1037
  assumes "lub_singleton lub"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1038
  shows "cont the_Sup (=) lub ord f"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1039
proof(rule contI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1040
  fix Y :: "'a set"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1041
  assume "Complete_Partial_Order.chain (=) Y" "Y \<noteq> {}"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1042
  then obtain a where "Y = {a}" by(auto simp add: chain_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1043
  thus "f (the_Sup Y) = lub (f ` Y)" using assms
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1044
    by(simp add: the_Sup_def lub_singleton_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1045
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1046
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1047
lemma mcont_eqI [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1048
  "\<lbrakk> class.preorder ord (mk_less ord); lub_singleton lub \<rbrakk>
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1049
  \<Longrightarrow> mcont the_Sup (=) lub ord f"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1050
  by(simp add: mcont_def cont_eqI monotone_eqI)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1051
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1052
subsection \<open>ccpo for products\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1053
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1054
definition prod_lub :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('b set \<Rightarrow> 'b) \<Rightarrow> ('a \<times> 'b) set \<Rightarrow> 'a \<times> 'b"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1055
  where "prod_lub Sup_a Sup_b Y = (Sup_a (fst ` Y), Sup_b (snd ` Y))"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1056
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1057
lemma lub_singleton_prod_lub [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1058
  "\<lbrakk> lub_singleton luba; lub_singleton lubb \<rbrakk> \<Longrightarrow> lub_singleton (prod_lub luba lubb)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1059
  by(simp add: lub_singleton_def prod_lub_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1060
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1061
lemma prod_lub_empty [simp]: "prod_lub luba lubb {} = (luba {}, lubb {})"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1062
  by(simp add: prod_lub_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1063
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1064
lemma preorder_rel_prodI [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1065
  assumes "class.preorder orda (mk_less orda)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1066
    and "class.preorder ordb (mk_less ordb)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1067
  shows "class.preorder (rel_prod orda ordb) (mk_less (rel_prod orda ordb))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1068
proof -
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1069
  interpret a: preorder orda "mk_less orda" by fact
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1070
  interpret b: preorder ordb "mk_less ordb" by fact
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1071
  show ?thesis by(unfold_locales)(auto simp add: mk_less_def intro: a.order_trans b.order_trans)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1072
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1073
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1074
lemma order_rel_prodI:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1075
  assumes a: "class.order orda (mk_less orda)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1076
    and b: "class.order ordb (mk_less ordb)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1077
  shows "class.order (rel_prod orda ordb) (mk_less (rel_prod orda ordb))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1078
    (is "class.order ?ord ?ord'")
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1079
proof(intro class.order.intro class.order_axioms.intro)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1080
  interpret a: order orda "mk_less orda" by(fact a)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1081
  interpret b: order ordb "mk_less ordb" by(fact b)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1082
  show "class.preorder ?ord ?ord'" by(rule preorder_rel_prodI) unfold_locales
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1083
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1084
  fix x y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1085
  assume "?ord x y" "?ord y x"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1086
  thus "x = y" by(cases x y rule: prod.exhaust[case_product prod.exhaust]) auto
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1087
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1088
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1089
lemma monotone_rel_prodI:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1090
  assumes mono2: "\<And>a. monotone ordb ordc (\<lambda>b. f (a, b))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1091
    and mono1: "\<And>b. monotone orda ordc (\<lambda>a. f (a, b))"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1092
    and a: "class.preorder orda (mk_less orda)"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1093
    and b: "class.preorder ordb (mk_less ordb)"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1094
    and c: "class.preorder ordc (mk_less ordc)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1095
  shows "monotone (rel_prod orda ordb) ordc f"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1096
proof -
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1097
  interpret a: preorder orda "mk_less orda" by(rule a)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1098
  interpret b: preorder ordb "mk_less ordb" by(rule b)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1099
  interpret c: preorder ordc "mk_less ordc" by(rule c)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1100
  show ?thesis using mono2 mono1
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1101
    by(auto 7 2 simp add: monotone_def intro: c.order_trans)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1102
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1103
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1104
lemma monotone_rel_prodD1:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1105
  assumes mono: "monotone (rel_prod orda ordb) ordc f"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1106
    and preorder: "class.preorder ordb (mk_less ordb)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1107
  shows "monotone orda ordc (\<lambda>a. f (a, b))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1108
proof -
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1109
  interpret preorder ordb "mk_less ordb" by(rule preorder)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1110
  show ?thesis using mono by(simp add: monotone_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1111
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1112
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1113
lemma monotone_rel_prodD2:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1114
  assumes mono: "monotone (rel_prod orda ordb) ordc f"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1115
    and preorder: "class.preorder orda (mk_less orda)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1116
  shows "monotone ordb ordc (\<lambda>b. f (a, b))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1117
proof -
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1118
  interpret preorder orda "mk_less orda" by(rule preorder)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1119
  show ?thesis using mono by(simp add: monotone_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1120
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1121
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1122
lemma monotone_case_prodI:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1123
  "\<lbrakk> \<And>a. monotone ordb ordc (f a); \<And>b. monotone orda ordc (\<lambda>a. f a b);
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1124
    class.preorder orda (mk_less orda); class.preorder ordb (mk_less ordb);
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1125
    class.preorder ordc (mk_less ordc) \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1126
  \<Longrightarrow> monotone (rel_prod orda ordb) ordc (case_prod f)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1127
  by(rule monotone_rel_prodI) simp_all
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1128
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1129
lemma monotone_case_prodD1:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1130
  assumes mono: "monotone (rel_prod orda ordb) ordc (case_prod f)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1131
    and preorder: "class.preorder ordb (mk_less ordb)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1132
  shows "monotone orda ordc (\<lambda>a. f a b)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1133
  using monotone_rel_prodD1[OF assms] by simp
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1134
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1135
lemma monotone_case_prodD2:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1136
  assumes mono: "monotone (rel_prod orda ordb) ordc (case_prod f)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1137
    and preorder: "class.preorder orda (mk_less orda)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1138
  shows "monotone ordb ordc (f a)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1139
  using monotone_rel_prodD2[OF assms] by simp
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1140
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1141
context 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1142
  fixes orda ordb ordc
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1143
  assumes a: "class.preorder orda (mk_less orda)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1144
    and b: "class.preorder ordb (mk_less ordb)"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1145
    and c: "class.preorder ordc (mk_less ordc)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1146
begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1147
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1148
lemma monotone_rel_prod_iff:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1149
  "monotone (rel_prod orda ordb) ordc f \<longleftrightarrow>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1150
   (\<forall>a. monotone ordb ordc (\<lambda>b. f (a, b))) \<and> 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1151
   (\<forall>b. monotone orda ordc (\<lambda>a. f (a, b)))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1152
  using a b c by(blast intro: monotone_rel_prodI dest: monotone_rel_prodD1 monotone_rel_prodD2)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1153
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1154
lemma monotone_case_prod_iff [simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1155
  "monotone (rel_prod orda ordb) ordc (case_prod f) \<longleftrightarrow>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1156
   (\<forall>a. monotone ordb ordc (f a)) \<and> (\<forall>b. monotone orda ordc (\<lambda>a. f a b))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1157
  by(simp add: monotone_rel_prod_iff)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1158
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1159
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1160
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1161
lemma monotone_case_prod_apply_iff:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1162
  "monotone orda ordb (\<lambda>x. (case_prod f x) y) \<longleftrightarrow> monotone orda ordb (case_prod (\<lambda>a b. f a b y))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1163
  by(simp add: monotone_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1164
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1165
lemma monotone_case_prod_applyD:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1166
  "monotone orda ordb (\<lambda>x. (case_prod f x) y)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1167
  \<Longrightarrow> monotone orda ordb (case_prod (\<lambda>a b. f a b y))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1168
  by(simp add: monotone_case_prod_apply_iff)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1169
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1170
lemma monotone_case_prod_applyI:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1171
  "monotone orda ordb (case_prod (\<lambda>a b. f a b y))
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1172
  \<Longrightarrow> monotone orda ordb (\<lambda>x. (case_prod f x) y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1173
  by(simp add: monotone_case_prod_apply_iff)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1174
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1175
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1176
lemma cont_case_prod_apply_iff:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1177
  "cont luba orda lubb ordb (\<lambda>x. (case_prod f x) y) \<longleftrightarrow> cont luba orda lubb ordb (case_prod (\<lambda>a b. f a b y))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1178
  by(simp add: cont_def split_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1179
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1180
lemma cont_case_prod_applyI:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1181
  "cont luba orda lubb ordb (case_prod (\<lambda>a b. f a b y))
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1182
  \<Longrightarrow> cont luba orda lubb ordb (\<lambda>x. (case_prod f x) y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1183
  by(simp add: cont_case_prod_apply_iff)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1184
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1185
lemma cont_case_prod_applyD:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1186
  "cont luba orda lubb ordb (\<lambda>x. (case_prod f x) y)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1187
  \<Longrightarrow> cont luba orda lubb ordb (case_prod (\<lambda>a b. f a b y))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1188
  by(simp add: cont_case_prod_apply_iff)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1189
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1190
lemma mcont_case_prod_apply_iff [simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1191
  "mcont luba orda lubb ordb (\<lambda>x. (case_prod f x) y) \<longleftrightarrow> 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1192
   mcont luba orda lubb ordb (case_prod (\<lambda>a b. f a b y))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1193
by(simp add: mcont_def monotone_case_prod_apply_iff cont_case_prod_apply_iff)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1194
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1195
lemma cont_prodD1: 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1196
  assumes cont: "cont (prod_lub luba lubb) (rel_prod orda ordb) lubc ordc f"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1197
  and "class.preorder orda (mk_less orda)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1198
  and luba: "lub_singleton luba"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1199
  shows "cont lubb ordb lubc ordc (\<lambda>y. f (x, y))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1200
proof(rule contI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1201
  interpret preorder orda "mk_less orda" by fact
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1202
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1203
  fix Y :: "'b set"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1204
  let ?Y = "{x} \<times> Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1205
  assume "Complete_Partial_Order.chain ordb Y" "Y \<noteq> {}"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1206
  hence "Complete_Partial_Order.chain (rel_prod orda ordb) ?Y" "?Y \<noteq> {}" 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1207
    by(simp_all add: chain_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1208
  with cont have "f (prod_lub luba lubb ?Y) = lubc (f ` ?Y)" by(rule contD)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1209
  moreover have "f ` ?Y = (\<lambda>y. f (x, y)) ` Y" by auto
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1210
  ultimately show "f (x, lubb Y) = lubc ((\<lambda>y. f (x, y)) ` Y)" using luba
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1211
    by(simp add: prod_lub_def \<open>Y \<noteq> {}\<close> lub_singleton_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1212
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1213
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1214
lemma cont_prodD2: 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1215
  assumes cont: "cont (prod_lub luba lubb) (rel_prod orda ordb) lubc ordc f"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1216
  and "class.preorder ordb (mk_less ordb)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1217
  and lubb: "lub_singleton lubb"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1218
  shows "cont luba orda lubc ordc (\<lambda>x. f (x, y))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1219
proof(rule contI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1220
  interpret preorder ordb "mk_less ordb" by fact
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1221
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1222
  fix Y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1223
  assume Y: "Complete_Partial_Order.chain orda Y" "Y \<noteq> {}"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1224
  let ?Y = "Y \<times> {y}"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1225
  have "f (luba Y, y) = f (prod_lub luba lubb ?Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1226
    using lubb by(simp add: prod_lub_def Y lub_singleton_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1227
  also from Y have "Complete_Partial_Order.chain (rel_prod orda ordb) ?Y" "?Y \<noteq> {}"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1228
    by(simp_all add: chain_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1229
  with cont have "f (prod_lub luba lubb ?Y) = lubc (f ` ?Y)" by(rule contD)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1230
  also have "f ` ?Y = (\<lambda>x. f (x, y)) ` Y" by auto
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1231
  finally show "f (luba Y, y) = lubc \<dots>" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1232
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1233
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1234
lemma cont_case_prodD1:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1235
  assumes "cont (prod_lub luba lubb) (rel_prod orda ordb) lubc ordc (case_prod f)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1236
  and "class.preorder orda (mk_less orda)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1237
  and "lub_singleton luba"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1238
  shows "cont lubb ordb lubc ordc (f x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1239
using cont_prodD1[OF assms] by simp
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1240
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1241
lemma cont_case_prodD2:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1242
  assumes "cont (prod_lub luba lubb) (rel_prod orda ordb) lubc ordc (case_prod f)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1243
  and "class.preorder ordb (mk_less ordb)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1244
  and "lub_singleton lubb"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1245
  shows "cont luba orda lubc ordc (\<lambda>x. f x y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1246
using cont_prodD2[OF assms] by simp
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1247
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1248
context ccpo begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1249
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1250
lemma cont_prodI: 
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1251
  assumes mono: "monotone (rel_prod orda ordb) (\<le>) f"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1252
  and cont1: "\<And>x. cont lubb ordb Sup (\<le>) (\<lambda>y. f (x, y))"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1253
  and cont2: "\<And>y. cont luba orda Sup (\<le>) (\<lambda>x. f (x, y))"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1254
  and "class.preorder orda (mk_less orda)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1255
  and "class.preorder ordb (mk_less ordb)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1256
  shows "cont (prod_lub luba lubb) (rel_prod orda ordb) Sup (\<le>) f"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1257
proof(rule contI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1258
  interpret a: preorder orda "mk_less orda" by fact 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1259
  interpret b: preorder ordb "mk_less ordb" by fact
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1260
  
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1261
  fix Y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1262
  assume chain: "Complete_Partial_Order.chain (rel_prod orda ordb) Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1263
    and "Y \<noteq> {}"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1264
  have "f (prod_lub luba lubb Y) = f (luba (fst ` Y), lubb (snd ` Y))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1265
    by(simp add: prod_lub_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1266
  also from cont2 have "f (luba (fst ` Y), lubb (snd ` Y)) = \<Squnion>((\<lambda>x. f (x, lubb (snd ` Y))) ` fst ` Y)"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1267
    by(rule contD)(simp_all add: chain_rel_prodD1[OF chain] \<open>Y \<noteq> {}\<close>)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1268
  also from cont1 have "\<And>x. f (x, lubb (snd ` Y)) = \<Squnion>((\<lambda>y. f (x, y)) ` snd ` Y)"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1269
    by(rule contD)(simp_all add: chain_rel_prodD2[OF chain] \<open>Y \<noteq> {}\<close>)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1270
  hence "\<Squnion>((\<lambda>x. f (x, lubb (snd ` Y))) ` fst ` Y) = \<Squnion>((\<lambda>x. \<dots> x) ` fst ` Y)" by simp
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1271
  also have "\<dots> = \<Squnion>((\<lambda>x. f (fst x, snd x)) ` Y)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1272
    unfolding image_image  using chain
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1273
  proof (rule diag_Sup)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1274
    show "\<And>y. y \<in> Y \<Longrightarrow> monotone (rel_prod orda ordb) (\<le>) (\<lambda>x. f (fst x, snd y))"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1275
      by (smt (verit, best) b.order_refl mono monotoneD monotoneI rel_prod_inject rel_prod_sel)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1276
  qed (use mono monotoneD in fastforce)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1277
  finally show "f (prod_lub luba lubb Y) = \<Squnion>(f ` Y)" by simp
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1278
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1279
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1280
lemma cont_case_prodI:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1281
  assumes "monotone (rel_prod orda ordb) (\<le>) (case_prod f)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1282
  and "\<And>x. cont lubb ordb Sup (\<le>) (\<lambda>y. f x y)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1283
  and "\<And>y. cont luba orda Sup (\<le>) (\<lambda>x. f x y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1284
  and "class.preorder orda (mk_less orda)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1285
  and "class.preorder ordb (mk_less ordb)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1286
  shows "cont (prod_lub luba lubb) (rel_prod orda ordb) Sup (\<le>) (case_prod f)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1287
by(rule cont_prodI)(simp_all add: assms)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1288
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1289
lemma cont_case_prod_iff:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1290
  "\<lbrakk> monotone (rel_prod orda ordb) (\<le>) (case_prod f);
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1291
     class.preorder orda (mk_less orda); lub_singleton luba;
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1292
     class.preorder ordb (mk_less ordb); lub_singleton lubb \<rbrakk>
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1293
  \<Longrightarrow> cont (prod_lub luba lubb) (rel_prod orda ordb) Sup (\<le>) (case_prod f) \<longleftrightarrow>
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1294
   (\<forall>x. cont lubb ordb Sup (\<le>) (\<lambda>y. f x y)) \<and> (\<forall>y. cont luba orda Sup (\<le>) (\<lambda>x. f x y))"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1295
by(blast dest: cont_case_prodD1 cont_case_prodD2 intro: cont_case_prodI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1296
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1297
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1298
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1299
context partial_function_definitions begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1300
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1301
lemma mono2mono2:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1302
  assumes f: "monotone (rel_prod ordb ordc) leq (\<lambda>(x, y). f x y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1303
  and t: "monotone orda ordb (\<lambda>x. t x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1304
  and t': "monotone orda ordc (\<lambda>x. t' x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1305
  shows "monotone orda leq (\<lambda>x. f (t x) (t' x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1306
  by (metis (mono_tags, lifting) case_prod_conv monotoneD monotoneI rel_prod.intros
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1307
      assms)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1308
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1309
lemma cont_case_prodI [cont_intro]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1310
  "\<lbrakk> monotone (rel_prod orda ordb) leq (case_prod f);
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1311
    \<And>x. cont lubb ordb lub leq (\<lambda>y. f x y);
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1312
    \<And>y. cont luba orda lub leq (\<lambda>x. f x y);
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1313
    class.preorder orda (mk_less orda);
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1314
    class.preorder ordb (mk_less ordb) \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1315
  \<Longrightarrow> cont (prod_lub luba lubb) (rel_prod orda ordb) lub leq (case_prod f)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1316
  by(rule ccpo.cont_case_prodI)(rule Partial_Function.ccpo[OF partial_function_definitions_axioms])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1317
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1318
lemma cont_case_prod_iff:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1319
  "\<lbrakk> monotone (rel_prod orda ordb) leq (case_prod f);
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1320
     class.preorder orda (mk_less orda); lub_singleton luba;
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1321
     class.preorder ordb (mk_less ordb); lub_singleton lubb \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1322
  \<Longrightarrow> cont (prod_lub luba lubb) (rel_prod orda ordb) lub leq (case_prod f) \<longleftrightarrow>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1323
   (\<forall>x. cont lubb ordb lub leq (\<lambda>y. f x y)) \<and> (\<forall>y. cont luba orda lub leq (\<lambda>x. f x y))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1324
  by(blast dest: cont_case_prodD1 cont_case_prodD2 intro: cont_case_prodI)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1325
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1326
lemma mcont_case_prod_iff [simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1327
  "\<lbrakk> class.preorder orda (mk_less orda); lub_singleton luba;
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1328
     class.preorder ordb (mk_less ordb); lub_singleton lubb \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1329
  \<Longrightarrow> mcont (prod_lub luba lubb) (rel_prod orda ordb) lub leq (case_prod f) \<longleftrightarrow>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1330
   (\<forall>x. mcont lubb ordb lub leq (\<lambda>y. f x y)) \<and> (\<forall>y. mcont luba orda lub leq (\<lambda>x. f x y))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1331
  unfolding mcont_def by(auto simp add: cont_case_prod_iff)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1332
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1333
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1334
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1335
lemma mono2mono_case_prod [cont_intro]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1336
  assumes "\<And>x y. monotone orda ordb (\<lambda>f. pair f x y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1337
  shows "monotone orda ordb (\<lambda>f. case_prod (pair f) x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1338
  by(rule monotoneI)(auto split: prod.split dest: monotoneD[OF assms])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1339
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1340
subsection \<open>Complete lattices as ccpo\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1341
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1342
context complete_lattice begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1343
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1344
lemma complete_lattice_ccpo: "class.ccpo Sup (\<le>) (<)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1345
  by(unfold_locales)(fast intro: Sup_upper Sup_least)+
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1346
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1347
lemma complete_lattice_ccpo': "class.ccpo Sup (\<le>) (mk_less (\<le>))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1348
  by(unfold_locales)(auto simp add: mk_less_def intro: Sup_upper Sup_least)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1349
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1350
lemma complete_lattice_partial_function_definitions: 
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1351
  "partial_function_definitions (\<le>) Sup"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1352
  by(unfold_locales)(auto intro: Sup_least Sup_upper)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1353
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1354
lemma complete_lattice_partial_function_definitions_dual:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1355
  "partial_function_definitions (\<ge>) Inf"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1356
  by(unfold_locales)(auto intro: Inf_lower Inf_greatest)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1357
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1358
lemmas [cont_intro, simp] =
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1359
  Partial_Function.ccpo[OF complete_lattice_partial_function_definitions]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1360
  Partial_Function.ccpo[OF complete_lattice_partial_function_definitions_dual]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1361
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1362
lemma mono2mono_inf:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1363
  assumes f: "monotone ord (\<le>) (\<lambda>x. f x)" 
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1364
    and g: "monotone ord (\<le>) (\<lambda>x. g x)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1365
  shows "monotone ord (\<le>) (\<lambda>x. f x \<sqinter> g x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1366
  by(auto 4 3 dest: monotoneD[OF f] monotoneD[OF g] intro: le_infI1 le_infI2 intro!: monotoneI)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1367
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1368
lemma mcont_const [simp]: "mcont lub ord Sup (\<le>) (\<lambda>_. c)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1369
  by(rule ccpo.mcont_const[OF complete_lattice_ccpo])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1370
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1371
lemma mono2mono_sup:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1372
  assumes f: "monotone ord (\<le>) (\<lambda>x. f x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1373
    and g: "monotone ord (\<le>) (\<lambda>x. g x)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1374
  shows "monotone ord (\<le>) (\<lambda>x. f x \<squnion> g x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1375
  by(auto 4 3 intro!: monotoneI intro: sup.coboundedI1 sup.coboundedI2 dest: monotoneD[OF f] monotoneD[OF g])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1376
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1377
lemma Sup_image_sup: 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1378
  assumes "Y \<noteq> {}"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1379
  shows "\<Squnion>((\<squnion>) x ` Y) = x \<squnion> \<Squnion>Y"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1380
proof(rule Sup_eqI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1381
  fix y
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1382
  assume "y \<in> (\<squnion>) x ` Y"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1383
  then obtain z where "y = x \<squnion> z" and "z \<in> Y" by blast
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1384
  from \<open>z \<in> Y\<close> have "z \<le> \<Squnion>Y" by(rule Sup_upper)
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1385
  with _ show "y \<le> x \<squnion> \<Squnion>Y" unfolding \<open>y = x \<squnion> z\<close> by(rule sup_mono) simp
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1386
next
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1387
  fix y
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1388
  assume upper: "\<And>z. z \<in> (\<squnion>) x ` Y \<Longrightarrow> z \<le> y"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1389
  show "x \<squnion> \<Squnion>Y \<le> y" unfolding Sup_insert[symmetric]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1390
  proof(rule Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1391
    fix z
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1392
    assume "z \<in> insert x Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1393
    from assms obtain z' where "z' \<in> Y" by blast
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1394
    let ?z = "if z \<in> Y then x \<squnion> z else x \<squnion> z'"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1395
    have "z \<le> x \<squnion> ?z" using \<open>z' \<in> Y\<close> \<open>z \<in> insert x Y\<close> by auto
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1396
    also have "\<dots> \<le> y" by(rule upper)(auto split: if_split_asm intro: \<open>z' \<in> Y\<close>)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1397
    finally show "z \<le> y" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1398
  qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1399
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1400
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1401
lemma mcont_sup1: "mcont Sup (\<le>) Sup (\<le>) (\<lambda>y. x \<squnion> y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1402
  by(auto 4 3 simp add: mcont_def sup.coboundedI1 sup.coboundedI2 intro!: monotoneI contI intro: Sup_image_sup[symmetric])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1403
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1404
lemma mcont_sup2: "mcont Sup (\<le>) Sup (\<le>) (\<lambda>x. x \<squnion> y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1405
  by(subst sup_commute)(rule mcont_sup1)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1406
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1407
lemma mcont2mcont_sup [cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1408
  "\<lbrakk> mcont lub ord Sup (\<le>) (\<lambda>x. f x);
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1409
     mcont lub ord Sup (\<le>) (\<lambda>x. g x) \<rbrakk>
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1410
  \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. f x \<squnion> g x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1411
  by(best intro: ccpo.mcont2mcont'[OF complete_lattice_ccpo] mcont_sup1 mcont_sup2 ccpo.mcont_const[OF complete_lattice_ccpo])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1412
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1413
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1414
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1415
lemmas [cont_intro] = admissible_leI[OF complete_lattice_ccpo']
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1416
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1417
context complete_distrib_lattice begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1418
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1419
lemma mcont_inf1: "mcont Sup (\<le>) Sup (\<le>) (\<lambda>y. x \<sqinter> y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1420
  by(auto intro: monotoneI contI simp add: le_infI2 inf_Sup mcont_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1421
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1422
lemma mcont_inf2: "mcont Sup (\<le>) Sup (\<le>) (\<lambda>x. x \<sqinter> y)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1423
  by(auto intro: monotoneI contI simp add: le_infI1 Sup_inf mcont_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1424
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1425
lemma mcont2mcont_inf [cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1426
  "\<lbrakk> mcont lub ord Sup (\<le>) (\<lambda>x. f x);
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1427
    mcont lub ord Sup (\<le>) (\<lambda>x. g x) \<rbrakk>
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1428
  \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. f x \<sqinter> g x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1429
  by(best intro: ccpo.mcont2mcont'[OF complete_lattice_ccpo] mcont_inf1 mcont_inf2 ccpo.mcont_const[OF complete_lattice_ccpo])
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1430
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1431
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1432
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1433
interpretation lfp: partial_function_definitions "(\<le>) :: _ :: complete_lattice \<Rightarrow> _" Sup
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1434
  by(rule complete_lattice_partial_function_definitions)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1435
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69546
diff changeset
  1436
declaration \<open>Partial_Function.init "lfp" \<^term>\<open>lfp.fixp_fun\<close> \<^term>\<open>lfp.mono_body\<close>
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1437
  @{thm lfp.fixp_rule_uc} @{thm lfp.fixp_induct_uc} NONE\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1438
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1439
interpretation gfp: partial_function_definitions "(\<ge>) :: _ :: complete_lattice \<Rightarrow> _" Inf
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1440
  by(rule complete_lattice_partial_function_definitions_dual)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1441
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69546
diff changeset
  1442
declaration \<open>Partial_Function.init "gfp" \<^term>\<open>gfp.fixp_fun\<close> \<^term>\<open>gfp.mono_body\<close>
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1443
  @{thm gfp.fixp_rule_uc} @{thm gfp.fixp_induct_uc} NONE\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1444
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1445
lemma insert_mono [partial_function_mono]:
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1446
  "monotone (fun_ord (\<subseteq>)) (\<subseteq>) A \<Longrightarrow> monotone (fun_ord (\<subseteq>)) (\<subseteq>) (\<lambda>y. insert x (A y))"
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1447
  by(rule monotoneI)(auto simp add: fun_ord_def dest: monotoneD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1448
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1449
lemma mono2mono_insert [THEN lfp.mono2mono, cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1450
  shows monotone_insert: "monotone (\<subseteq>) (\<subseteq>) (insert x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1451
  by(rule monotoneI) blast
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1452
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1453
lemma mcont2mcont_insert[THEN lfp.mcont2mcont, cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1454
  shows mcont_insert: "mcont Union (\<subseteq>) Union (\<subseteq>) (insert x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1455
  by(blast intro: mcontI contI monotone_insert)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1456
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1457
lemma mono2mono_image [THEN lfp.mono2mono, cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1458
  shows monotone_image: "monotone (\<subseteq>) (\<subseteq>) ((`) f)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1459
  by (simp add: image_mono monoI)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1460
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1461
lemma cont_image: "cont Union (\<subseteq>) Union (\<subseteq>) ((`) f)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1462
  by (meson contI image_Union)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1463
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1464
lemma mcont2mcont_image [THEN lfp.mcont2mcont, cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1465
  shows mcont_image: "mcont Union (\<subseteq>) Union (\<subseteq>) ((`) f)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1466
  by(blast intro: mcontI monotone_image cont_image)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1467
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1468
context complete_lattice begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1469
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1470
lemma monotone_Sup [cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1471
  "monotone ord (\<subseteq>) f \<Longrightarrow> monotone ord (\<le>) (\<lambda>x. \<Squnion>f x)"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1472
  by(blast intro: monotoneI Sup_least Sup_upper dest: monotoneD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1473
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1474
lemma cont_Sup:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1475
  assumes "cont lub ord Union (\<subseteq>) f"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1476
  shows "cont lub ord Sup (\<le>) (\<lambda>x. \<Squnion>f x)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1477
proof -
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1478
  have "\<And>Y. \<lbrakk>Complete_Partial_Order.chain ord Y; Y \<noteq> {}\<rbrakk>
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1479
         \<Longrightarrow> \<Squnion> \<Union> (f ` Y) = (\<Squnion>x\<in>Y. \<Squnion> f x)"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1480
    by (blast intro: Sup_least Sup_upper order_trans order.antisym)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1481
  with assms show ?thesis
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1482
    by (force simp: cont_def)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1483
qed
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1484
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1485
lemma mcont_Sup: "mcont lub ord Union (\<subseteq>) f \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. \<Squnion>f x)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1486
  unfolding mcont_def by(blast intro: monotone_Sup cont_Sup)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1487
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1488
lemma monotone_SUP:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1489
  "\<lbrakk> monotone ord (\<subseteq>) f; \<And>y. monotone ord (\<le>) (\<lambda>x. g x y) \<rbrakk> \<Longrightarrow> monotone ord (\<le>) (\<lambda>x. \<Squnion>y\<in>f x. g x y)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1490
  by(rule monotoneI)(blast dest: monotoneD intro: Sup_upper order_trans intro!: Sup_least)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1491
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1492
lemma monotone_SUP2:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1493
  "(\<And>y. y \<in> A \<Longrightarrow> monotone ord (\<le>) (\<lambda>x. g x y)) \<Longrightarrow> monotone ord (\<le>) (\<lambda>x. \<Squnion>y\<in>A. g x y)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1494
  by(rule monotoneI)(blast intro: Sup_upper order_trans dest: monotoneD intro!: Sup_least)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1495
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1496
lemma cont_SUP:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1497
  assumes f: "mcont lub ord Union (\<subseteq>) f"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1498
  and g: "\<And>y. mcont lub ord Sup (\<le>) (\<lambda>x. g x y)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1499
  shows "cont lub ord Sup (\<le>) (\<lambda>x. \<Squnion>y\<in>f x. g x y)"
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1500
proof(rule contI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1501
  fix Y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1502
  assume chain: "Complete_Partial_Order.chain ord Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1503
    and Y: "Y \<noteq> {}"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1504
  show "\<Squnion>(g (lub Y) ` f (lub Y)) = \<Squnion>((\<lambda>x. \<Squnion>(g x ` f x)) ` Y)" (is "?lhs = ?rhs")
73411
1f1366966296 avoid name clash
haftmann
parents: 70961
diff changeset
  1505
  proof(rule order.antisym)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1506
    show "?lhs \<le> ?rhs"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1507
    proof(rule Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1508
      fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1509
      assume "x \<in> g (lub Y) ` f (lub Y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1510
      with mcont_contD[OF f chain Y] mcont_contD[OF g chain Y]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1511
      obtain y z where "y \<in> Y" "z \<in> f y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1512
        and x: "x = \<Squnion>((\<lambda>x. g x z) ` Y)" by auto
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1513
      show "x \<le> ?rhs" unfolding x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1514
      proof(rule Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1515
        fix u
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1516
        assume "u \<in> (\<lambda>x. g x z) ` Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1517
        then obtain y' where "u = g y' z" "y' \<in> Y" by auto
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1518
        from chain \<open>y \<in> Y\<close> \<open>y' \<in> Y\<close> have "ord y y' \<or> ord y' y" by(rule chainD)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1519
        thus "u \<le> ?rhs"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1520
        proof
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1521
          note \<open>u = g y' z\<close> also
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1522
          assume "ord y y'"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1523
          with f have "f y \<subseteq> f y'" by(rule mcont_monoD)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1524
          with \<open>z \<in> f y\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1525
          have "g y' z \<le> \<Squnion>(g y' ` f y')" by(auto intro: Sup_upper)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1526
          also have "\<dots> \<le> ?rhs" using \<open>y' \<in> Y\<close> by(auto intro: Sup_upper)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1527
          finally show ?thesis .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1528
        next
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1529
          note \<open>u = g y' z\<close> also
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1530
          assume "ord y' y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1531
          with g have "g y' z \<le> g y z" by(rule mcont_monoD)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1532
          also have "\<dots> \<le> \<Squnion>(g y ` f y)" using \<open>z \<in> f y\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1533
            by(auto intro: Sup_upper)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1534
          also have "\<dots> \<le> ?rhs" using \<open>y \<in> Y\<close> by(auto intro: Sup_upper)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1535
          finally show ?thesis .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1536
        qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1537
      qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1538
    qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1539
  next
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1540
    show "?rhs \<le> ?lhs"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1541
    proof(rule Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1542
      fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1543
      assume "x \<in> (\<lambda>x. \<Squnion>(g x ` f x)) ` Y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1544
      then obtain y where x: "x = \<Squnion>(g y ` f y)" and "y \<in> Y" by auto
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1545
      show "x \<le> ?lhs" unfolding x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1546
      proof(rule Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1547
        fix u
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1548
        assume "u \<in> g y ` f y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1549
        then obtain z where "u = g y z" "z \<in> f y" by auto
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1550
        note \<open>u = g y z\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1551
        also have "g y z \<le> \<Squnion>((\<lambda>x. g x z) ` Y)"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1552
          using \<open>y \<in> Y\<close> by(auto intro: Sup_upper)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1553
        also have "\<dots> = g (lub Y) z" by(simp add: mcont_contD[OF g chain Y])
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1554
        also have "\<dots> \<le> ?lhs" using \<open>z \<in> f y\<close> \<open>y \<in> Y\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1555
          by(auto intro: Sup_upper simp add: mcont_contD[OF f chain Y])
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1556
        finally show "u \<le> ?lhs" .
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1557
      qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1558
    qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1559
  qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1560
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1561
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1562
lemma mcont_SUP [cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1563
  "\<lbrakk> mcont lub ord Union (\<subseteq>) f; \<And>y. mcont lub ord Sup (\<le>) (\<lambda>x. g x y) \<rbrakk>
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1564
  \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. \<Squnion>y\<in>f x. g x y)"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  1565
by(blast intro: mcontI cont_SUP monotone_SUP mcont_mono)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1566
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1567
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1568
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1569
lemma admissible_Ball [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1570
  "\<lbrakk> \<And>x. ccpo.admissible lub ord (\<lambda>A. P A x);
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1571
     mcont lub ord Union (\<subseteq>) f;
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1572
     class.ccpo lub ord (mk_less ord) \<rbrakk>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1573
  \<Longrightarrow> ccpo.admissible lub ord (\<lambda>A. \<forall>x\<in>f A. P A x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1574
unfolding Ball_def by simp
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1575
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1576
lemma admissible_Bex'[THEN admissible_subst, cont_intro, simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66244
diff changeset
  1577
  shows admissible_Bex: "ccpo.admissible Union (\<subseteq>) (\<lambda>A. \<exists>x\<in>A. P x)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1578
  using ccpo.admissible_def by fastforce
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1579
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62652
diff changeset
  1580
subsection \<open>Parallel fixpoint induction\<close>
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1581
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1582
context
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1583
  fixes luba :: "'a set \<Rightarrow> 'a"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1584
  and orda :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1585
  and lubb :: "'b set \<Rightarrow> 'b"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1586
  and ordb :: "'b \<Rightarrow> 'b \<Rightarrow> bool"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1587
  assumes a: "class.ccpo luba orda (mk_less orda)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1588
  and b: "class.ccpo lubb ordb (mk_less ordb)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1589
begin
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1590
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1591
interpretation a: ccpo luba orda "mk_less orda" by(rule a)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1592
interpretation b: ccpo lubb ordb "mk_less ordb" by(rule b)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1593
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1594
lemma ccpo_rel_prodI:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1595
  "class.ccpo (prod_lub luba lubb) (rel_prod orda ordb) (mk_less (rel_prod orda ordb))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1596
  (is "class.ccpo ?lub ?ord ?ord'")
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1597
proof(intro class.ccpo.intro class.ccpo_axioms.intro)
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1598
  show "class.order ?ord ?ord'" 
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1599
    by(rule order_rel_prodI) intro_locales
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1600
  show "\<And>A x. \<lbrakk>Complete_Partial_Order.chain (rel_prod orda ordb) A; x \<in> A\<rbrakk>
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1601
           \<Longrightarrow> rel_prod orda ordb x (prod_lub luba lubb A)"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1602
    by (simp add: a.ccpo_Sup_upper b.ccpo_Sup_upper chain_rel_prodD1 chain_rel_prodD2
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1603
        prod_lub_def rel_prod_sel)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1604
  show "\<And>A z. \<lbrakk>Complete_Partial_Order.chain (rel_prod orda ordb) A;
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1605
            \<And>x. x \<in> A \<Longrightarrow> rel_prod orda ordb x z\<rbrakk>
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1606
           \<Longrightarrow> rel_prod orda ordb (prod_lub luba lubb A) z"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1607
    by (metis (full_types) a.ccpo_Sup_below_iff b.ccpo_Sup_least chain_rel_prodD1
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1608
        chain_rel_prodD2 imageE prod.sel prod_lub_def rel_prod_sel)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1609
qed
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1610
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1611
interpretation ab: ccpo "prod_lub luba lubb" "rel_prod orda ordb" "mk_less (rel_prod orda ordb)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1612
by(rule ccpo_rel_prodI)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1613
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1614
lemma monotone_map_prod [simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1615
  "monotone (rel_prod orda ordb) (rel_prod ordc ordd) (map_prod f g) \<longleftrightarrow>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1616
   monotone orda ordc f \<and> monotone ordb ordd g"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1617
by(auto simp add: monotone_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1618
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1619
lemma parallel_fixp_induct:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1620
  assumes adm: "ccpo.admissible (prod_lub luba lubb) (rel_prod orda ordb) (\<lambda>x. P (fst x) (snd x))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1621
  and f: "monotone orda orda f"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1622
  and g: "monotone ordb ordb g"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1623
  and bot: "P (luba {}) (lubb {})"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1624
  and step: "\<And>x y. P x y \<Longrightarrow> P (f x) (g y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1625
  shows "P (ccpo.fixp luba orda f) (ccpo.fixp lubb ordb g)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1626
proof -
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1627
  let ?lub = "prod_lub luba lubb"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1628
    and ?ord = "rel_prod orda ordb"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1629
    and ?P = "\<lambda>(x, y). P x y"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1630
  from adm have adm': "ccpo.admissible ?lub ?ord ?P" by(simp add: split_def)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1631
  hence "?P (ccpo.fixp (prod_lub luba lubb) (rel_prod orda ordb) (map_prod f g))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1632
    by(rule ab.fixp_induct)(auto simp add: f g step bot)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1633
  also have "ccpo.fixp (prod_lub luba lubb) (rel_prod orda ordb) (map_prod f g) = 
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1634
            (ccpo.fixp luba orda f, ccpo.fixp lubb ordb g)" (is "?lhs = (?rhs1, ?rhs2)")
73411
1f1366966296 avoid name clash
haftmann
parents: 70961
diff changeset
  1635
  proof(rule ab.order.antisym)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1636
    have "ccpo.admissible ?lub ?ord (\<lambda>xy. ?ord xy (?rhs1, ?rhs2))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1637
      by(rule admissible_leI[OF ccpo_rel_prodI])(auto simp add: prod_lub_def chain_empty intro: a.ccpo_Sup_least b.ccpo_Sup_least)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1638
    thus "?ord ?lhs (?rhs1, ?rhs2)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1639
      by(rule ab.fixp_induct)(auto 4 3 dest: monotoneD[OF f] monotoneD[OF g] simp add: b.fixp_unfold[OF g, symmetric] a.fixp_unfold[OF f, symmetric] f g intro: a.ccpo_Sup_least b.ccpo_Sup_least chain_empty)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1640
  next
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1641
    have "ccpo.admissible luba orda (\<lambda>x. orda x (fst ?lhs))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1642
      by(rule admissible_leI[OF a])(auto intro: a.ccpo_Sup_least simp add: chain_empty)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1643
    hence "orda ?rhs1 (fst ?lhs)" using f
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1644
    proof(rule a.fixp_induct)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1645
      fix x
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1646
      assume "orda x (fst ?lhs)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1647
      thus "orda (f x) (fst ?lhs)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1648
        by(subst ab.fixp_unfold)(auto simp add: f g dest: monotoneD[OF f])
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1649
    qed(auto intro: a.ccpo_Sup_least chain_empty)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1650
    moreover
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1651
    have "ccpo.admissible lubb ordb (\<lambda>y. ordb y (snd ?lhs))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1652
      by(rule admissible_leI[OF b])(auto intro: b.ccpo_Sup_least simp add: chain_empty)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1653
    hence "ordb ?rhs2 (snd ?lhs)" using g
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1654
    proof(rule b.fixp_induct)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1655
      fix y
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1656
      assume "ordb y (snd ?lhs)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1657
      thus "ordb (g y) (snd ?lhs)"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1658
        by (smt (verit, best) ab.fixp_unfold f g monotoneD monotone_map_prod
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1659
            snd_map_prod)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1660
    qed(auto intro: b.ccpo_Sup_least chain_empty)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1661
    ultimately show "?ord (?rhs1, ?rhs2) ?lhs"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1662
      by(simp add: rel_prod_conv split_beta)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1663
  qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1664
  finally show ?thesis by simp
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1665
qed
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1666
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1667
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1668
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1669
lemma parallel_fixp_induct_uc:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1670
  assumes a: "partial_function_definitions orda luba"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1671
  and b: "partial_function_definitions ordb lubb"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1672
  and F: "\<And>x. monotone (fun_ord orda) orda (\<lambda>f. U1 (F (C1 f)) x)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1673
  and G: "\<And>y. monotone (fun_ord ordb) ordb (\<lambda>g. U2 (G (C2 g)) y)"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1674
  and eq1: "f \<equiv> C1 (ccpo.fixp (fun_lub luba) (fun_ord orda) (\<lambda>f. U1 (F (C1 f))))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1675
  and eq2: "g \<equiv> C2 (ccpo.fixp (fun_lub lubb) (fun_ord ordb) (\<lambda>g. U2 (G (C2 g))))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1676
  and inverse: "\<And>f. U1 (C1 f) = f"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1677
  and inverse2: "\<And>g. U2 (C2 g) = g"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1678
  and adm: "ccpo.admissible (prod_lub (fun_lub luba) (fun_lub lubb)) (rel_prod (fun_ord orda) (fun_ord ordb)) (\<lambda>x. P (fst x) (snd x))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1679
  and bot: "P (\<lambda>_. luba {}) (\<lambda>_. lubb {})"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1680
  and step: "\<And>f g. P (U1 f) (U2 g) \<Longrightarrow> P (U1 (F f)) (U2 (G g))"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1681
shows "P (U1 f) (U2 g)"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1682
  unfolding eq1 eq2 inverse inverse2
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1683
proof (rule parallel_fixp_induct[OF partial_function_definitions.ccpo[OF a] partial_function_definitions.ccpo[OF b] adm])
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1684
  show "monotone (fun_ord orda) (fun_ord orda) (\<lambda>f. U1 (F (C1 f)))"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1685
    "monotone (fun_ord ordb) (fun_ord ordb) (\<lambda>g. U2 (G (C2 g)))"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1686
    using F G by(simp_all add: monotone_def fun_ord_def)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1687
  show "P (fun_lub luba {}) (fun_lub lubb {})"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1688
    by (simp add: fun_lub_def bot)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1689
  show "\<And>x y. P x y \<Longrightarrow> P (U1 (F (C1 x))) (U2 (G (C2 y)))"
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1690
    by (simp add: inverse inverse2 local.step)
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1691
qed
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1692
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1693
lemmas parallel_fixp_induct_1_1 = parallel_fixp_induct_uc[
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1694
  of _ _ _ _ "\<lambda>x. x" _ "\<lambda>x. x" "\<lambda>x. x" _ "\<lambda>x. x",
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1695
  OF _ _ _ _ _ _ refl refl]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1696
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1697
lemmas parallel_fixp_induct_2_2 = parallel_fixp_induct_uc[
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1698
  of _ _ _ _ "case_prod" _ "curry" "case_prod" _ "curry",
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1699
  where P="\<lambda>f g. P (curry f) (curry g)",
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1700
  unfolded case_prod_curry curry_case_prod curry_K,
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1701
  OF _ _ _ _ _ _ refl refl]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1702
  for P
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1703
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1704
lemma monotone_fst: "monotone (rel_prod orda ordb) orda fst"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1705
  by(auto intro: monotoneI)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1706
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1707
lemma mcont_fst: "mcont (prod_lub luba lubb) (rel_prod orda ordb) luba orda fst"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1708
  by(auto intro!: mcontI monotoneI contI simp add: prod_lub_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1709
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1710
lemma mcont2mcont_fst [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1711
  "mcont lub ord (prod_lub luba lubb) (rel_prod orda ordb) t
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1712
  \<Longrightarrow> mcont lub ord luba orda (\<lambda>x. fst (t x))"
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1713
  by (simp add: mcont_def monotone_on_def prod_lub_def cont_def image_image rel_prod_sel)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1714
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1715
lemma monotone_snd: "monotone (rel_prod orda ordb) ordb snd"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1716
  by(auto intro: monotoneI)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1717
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1718
lemma mcont_snd: "mcont (prod_lub luba lubb) (rel_prod orda ordb) lubb ordb snd"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1719
  by(auto intro!: mcontI monotoneI contI simp add: prod_lub_def)
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1720
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1721
lemma mcont2mcont_snd [cont_intro, simp]:
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1722
  "mcont lub ord (prod_lub luba lubb) (rel_prod orda ordb) t
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1723
  \<Longrightarrow> mcont lub ord lubb ordb (\<lambda>x. snd (t x))"
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1724
by(auto intro!: mcontI monotoneI contI dest: mcont_monoD mcont_contD simp add: rel_prod_sel split_beta prod_lub_def image_image)
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1725
63243
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1726
lemma monotone_Pair:
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1727
  "\<lbrakk> monotone ord orda f; monotone ord ordb g \<rbrakk>
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1728
  \<Longrightarrow> monotone ord (rel_prod orda ordb) (\<lambda>x. (f x, g x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1729
  by(simp add: monotone_def)
63243
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1730
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1731
lemma cont_Pair:
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1732
  "\<lbrakk> cont lub ord luba orda f; cont lub ord lubb ordb g \<rbrakk>
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1733
  \<Longrightarrow> cont lub ord (prod_lub luba lubb) (rel_prod orda ordb) (\<lambda>x. (f x, g x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1734
  by(rule contI)(auto simp add: prod_lub_def image_image dest!: contD)
63243
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1735
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1736
lemma mcont_Pair:
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1737
  "\<lbrakk> mcont lub ord luba orda f; mcont lub ord lubb ordb g \<rbrakk>
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1738
  \<Longrightarrow> mcont lub ord (prod_lub luba lubb) (rel_prod orda ordb) (\<lambda>x. (f x, g x))"
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1739
  by(rule mcontI)(simp_all add: monotone_Pair mcont_mono cont_Pair)
63243
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1740
81806
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1741
context partial_function_definitions 
602639414559 Some work on an ancient theory file. And a weird failure in Float.thy
paulson <lp15@cam.ac.uk>
parents: 81281
diff changeset
  1742
begin
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1743
text \<open>Specialised versions of @{thm [source] mcont_call} for admissibility proofs for parallel fixpoint inductions\<close>
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1744
lemmas mcont_call_fst [cont_intro] = mcont_call[THEN mcont2mcont, OF mcont_fst]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1745
lemmas mcont_call_snd [cont_intro] = mcont_call[THEN mcont2mcont, OF mcont_snd]
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1746
end
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1747
63243
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1748
lemma map_option_mono [partial_function_mono]:
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1749
  "mono_option B \<Longrightarrow> mono_option (\<lambda>f. map_option g (B f))"
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1750
unfolding map_conv_bind_option by(rule bind_mono) simp_all
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1751
66244
4c999b5d78e2 qualify Complete_Partial_Order2.compact
Andreas Lochbihler
parents: 65366
diff changeset
  1752
lemma compact_flat_lub [cont_intro]: "ccpo.compact (flat_lub x) (flat_ord x) y"
63243
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1753
using flat_interpretation[THEN ccpo]
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1754
proof(rule ccpo.compactI[OF _ ccpo.admissibleI])
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1755
  fix A
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1756
  assume chain: "Complete_Partial_Order.chain (flat_ord x) A"
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1757
    and A: "A \<noteq> {}"
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1758
    and *: "\<forall>z\<in>A. \<not> flat_ord x y z"
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1759
  from A obtain z where "z \<in> A" by blast
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1760
  with * have z: "\<not> flat_ord x y z" ..
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1761
  hence y: "x \<noteq> y" "y \<noteq> z" by(auto simp add: flat_ord_def)
81281
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
  1762
  have "y \<noteq> (THE z. z \<in> A - {x})" if "\<not> A \<subseteq> {x}"
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
  1763
  proof -
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
  1764
    from that obtain z' where "z' \<in> A" "z' \<noteq> x" by auto
63243
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1765
    then have "(THE z. z \<in> A - {x}) = z'"
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1766
      by(intro the_equality)(auto dest: chainD[OF chain] simp add: flat_ord_def)
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1767
    moreover have "z' \<noteq> y" using \<open>z' \<in> A\<close> * by(auto simp add: flat_ord_def)
81281
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
  1768
    ultimately show ?thesis by simp
c1e418161ace tuned proofs;
wenzelm
parents: 80914
diff changeset
  1769
  qed
81982
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1770
  with z show "\<not> flat_ord x y (flat_lub x A)" 
paulson <lp15@cam.ac.uk>
parents: 81806
diff changeset
  1771
    by(simp add: flat_ord_def flat_lub_def)
63243
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1772
qed
1bc6816fd525 add theory of discrete subprobability distributions
Andreas Lochbihler
parents: 63170
diff changeset
  1773
62652
7248d106c607 move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
diff changeset
  1774
end