src/HOL/HOLCF/Deflation.thy
author wenzelm
Sun, 03 Sep 2023 19:28:59 +0200
changeset 78643 d5a1d64a563d
parent 67312 0d25e02759b7
child 81577 a712bf5ccab0
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41430
diff changeset
     1
(*  Title:      HOL/HOLCF/Deflation.thy
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
     3
*)
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
     4
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
     5
section \<open>Continuous deflations and ep-pairs\<close>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
     6
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
     7
theory Deflation
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
     8
  imports Cfun
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
     9
begin
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    10
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35901
diff changeset
    11
default_sort cpo
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    12
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    13
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
    14
subsection \<open>Continuous deflations\<close>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    15
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    16
locale deflation =
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    17
  fixes d :: "'a \<rightarrow> 'a"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    18
  assumes idem: "\<And>x. d\<cdot>(d\<cdot>x) = d\<cdot>x"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
    19
  assumes below: "\<And>x. d\<cdot>x \<sqsubseteq> x"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    20
begin
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    21
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
    22
lemma below_ID: "d \<sqsubseteq> ID"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    23
  by (rule cfun_belowI) (simp add: below)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    24
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
    25
text \<open>The set of fixed points is the same as the range.\<close>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    26
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    27
lemma fixes_eq_range: "{x. d\<cdot>x = x} = range (\<lambda>x. d\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    28
  by (auto simp add: eq_sym_conv idem)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    29
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    30
lemma range_eq_fixes: "range (\<lambda>x. d\<cdot>x) = {x. d\<cdot>x = x}"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    31
  by (auto simp add: eq_sym_conv idem)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    32
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
    33
text \<open>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    34
  The pointwise ordering on deflation functions coincides with
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    35
  the subset ordering of their sets of fixed-points.
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
    36
\<close>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    37
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
    38
lemma belowI:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    39
  assumes f: "\<And>x. d\<cdot>x = x \<Longrightarrow> f\<cdot>x = x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    40
  shows "d \<sqsubseteq> f"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 39985
diff changeset
    41
proof (rule cfun_belowI)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    42
  fix x
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    43
  from below have "f\<cdot>(d\<cdot>x) \<sqsubseteq> f\<cdot>x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    44
    by (rule monofun_cfun_arg)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    45
  also from idem have "f\<cdot>(d\<cdot>x) = d\<cdot>x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    46
    by (rule f)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    47
  finally show "d\<cdot>x \<sqsubseteq> f\<cdot>x" .
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    48
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    49
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
    50
lemma belowD: "\<lbrakk>f \<sqsubseteq> d; f\<cdot>x = x\<rbrakk> \<Longrightarrow> d\<cdot>x = x"
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
    51
proof (rule below_antisym)
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
    52
  from below show "d\<cdot>x \<sqsubseteq> x" .
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    53
  assume "f \<sqsubseteq> d"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    54
  then have "f\<cdot>x \<sqsubseteq> d\<cdot>x" by (rule monofun_cfun_fun)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    55
  also assume "f\<cdot>x = x"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    56
  finally show "x \<sqsubseteq> d\<cdot>x" .
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    57
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    58
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    59
end
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    60
33503
3496616b2171 lemma deflation_strict
huffman
parents: 31076
diff changeset
    61
lemma deflation_strict: "deflation d \<Longrightarrow> d\<cdot>\<bottom> = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    62
  by (rule deflation.below [THEN bottomI])
33503
3496616b2171 lemma deflation_strict
huffman
parents: 31076
diff changeset
    63
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    64
lemma adm_deflation: "adm (\<lambda>d. deflation d)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    65
  by (simp add: deflation_def)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    66
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    67
lemma deflation_ID: "deflation ID"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    68
  by (simp add: deflation.intro)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    69
41430
1aa23e9f2c87 change some lemma names containing 'UU' to 'bottom'
huffman
parents: 41182
diff changeset
    70
lemma deflation_bottom: "deflation \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    71
  by (simp add: deflation.intro)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    72
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    73
lemma deflation_below_iff: "deflation p \<Longrightarrow> deflation q \<Longrightarrow> p \<sqsubseteq> q \<longleftrightarrow> (\<forall>x. p\<cdot>x = x \<longrightarrow> q\<cdot>x = x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    74
  apply safe
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    75
   apply (simp add: deflation.belowD)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    76
  apply (simp add: deflation.belowI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    77
  done
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    78
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
    79
text \<open>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    80
  The composition of two deflations is equal to
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    81
  the lesser of the two (if they are comparable).
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
    82
\<close>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    83
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
    84
lemma deflation_below_comp1:
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
    85
  assumes "deflation f"
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
    86
  assumes "deflation g"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    87
  shows "f \<sqsubseteq> g \<Longrightarrow> f\<cdot>(g\<cdot>x) = f\<cdot>x"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
    88
proof (rule below_antisym)
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28613
diff changeset
    89
  interpret g: deflation g by fact
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
    90
  from g.below show "f\<cdot>(g\<cdot>x) \<sqsubseteq> f\<cdot>x" by (rule monofun_cfun_arg)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    91
next
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28613
diff changeset
    92
  interpret f: deflation f by fact
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    93
  assume "f \<sqsubseteq> g"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    94
  then have "f\<cdot>x \<sqsubseteq> g\<cdot>x" by (rule monofun_cfun_fun)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
    95
  then have "f\<cdot>(f\<cdot>x) \<sqsubseteq> f\<cdot>(g\<cdot>x)" by (rule monofun_cfun_arg)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    96
  also have "f\<cdot>(f\<cdot>x) = f\<cdot>x" by (rule f.idem)
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    97
  finally show "f\<cdot>x \<sqsubseteq> f\<cdot>(g\<cdot>x)" .
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    98
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
    99
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   100
lemma deflation_below_comp2: "deflation f \<Longrightarrow> deflation g \<Longrightarrow> f \<sqsubseteq> g \<Longrightarrow> g\<cdot>(f\<cdot>x) = f\<cdot>x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   101
  by (simp only: deflation.belowD deflation.idem)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   102
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   103
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   104
subsection \<open>Deflations with finite range\<close>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   105
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   106
lemma finite_range_imp_finite_fixes:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   107
  assumes "finite (range f)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   108
  shows "finite {x. f x = x}"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   109
proof -
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   110
  have "{x. f x = x} \<subseteq> range f"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   111
    by (clarify, erule subst, rule rangeI)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   112
  from this assms show "finite {x. f x = x}"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   113
    by (rule finite_subset)
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   114
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   115
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   116
locale finite_deflation = deflation +
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   117
  assumes finite_fixes: "finite {x. d\<cdot>x = x}"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   118
begin
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   119
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   120
lemma finite_range: "finite (range (\<lambda>x. d\<cdot>x))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   121
  by (simp add: range_eq_fixes finite_fixes)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   122
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   123
lemma finite_image: "finite ((\<lambda>x. d\<cdot>x) ` A)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   124
  by (rule finite_subset [OF image_mono [OF subset_UNIV] finite_range])
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   125
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   126
lemma compact: "compact (d\<cdot>x)"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   127
proof (rule compactI2)
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   128
  fix Y :: "nat \<Rightarrow> 'a"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   129
  assume Y: "chain Y"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   130
  have "finite_chain (\<lambda>i. d\<cdot>(Y i))"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   131
  proof (rule finite_range_imp_finch)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   132
    from Y show "chain (\<lambda>i. d\<cdot>(Y i))" by simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   133
    have "range (\<lambda>i. d\<cdot>(Y i)) \<subseteq> range (\<lambda>x. d\<cdot>x)" by auto
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   134
    then show "finite (range (\<lambda>i. d\<cdot>(Y i)))"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   135
      using finite_range by (rule finite_subset)
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   136
  qed
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   137
  then have "\<exists>j. (\<Squnion>i. d\<cdot>(Y i)) = d\<cdot>(Y j)"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   138
    by (simp add: finite_chain_def maxinch_is_thelub Y)
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   139
  then obtain j where j: "(\<Squnion>i. d\<cdot>(Y i)) = d\<cdot>(Y j)" ..
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   140
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   141
  assume "d\<cdot>x \<sqsubseteq> (\<Squnion>i. Y i)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   142
  then have "d\<cdot>(d\<cdot>x) \<sqsubseteq> d\<cdot>(\<Squnion>i. Y i)"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   143
    by (rule monofun_cfun_arg)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   144
  then have "d\<cdot>x \<sqsubseteq> (\<Squnion>i. d\<cdot>(Y i))"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   145
    by (simp add: contlub_cfun_arg Y idem)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   146
  with j have "d\<cdot>x \<sqsubseteq> d\<cdot>(Y j)" by simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   147
  then have "d\<cdot>x \<sqsubseteq> Y j"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   148
    using below by (rule below_trans)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   149
  then show "\<exists>j. d\<cdot>x \<sqsubseteq> Y j" ..
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   150
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   151
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   152
end
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   153
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   154
lemma finite_deflation_intro: "deflation d \<Longrightarrow> finite {x. d\<cdot>x = x} \<Longrightarrow> finite_deflation d"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   155
  by (intro finite_deflation.intro finite_deflation_axioms.intro)
39973
c62b4ff97bfc add lemma finite_deflation_intro
Brian Huffman <brianh@cs.pdx.edu>
parents: 39971
diff changeset
   156
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   157
lemma finite_deflation_imp_deflation: "finite_deflation d \<Longrightarrow> deflation d"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   158
  by (simp add: finite_deflation_def)
39971
2949af5e6b9c move lemmas to Deflation.thy
Brian Huffman <brianh@cs.pdx.edu>
parents: 36452
diff changeset
   159
41430
1aa23e9f2c87 change some lemma names containing 'UU' to 'bottom'
huffman
parents: 41182
diff changeset
   160
lemma finite_deflation_bottom: "finite_deflation \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   161
  by standard simp_all
39971
2949af5e6b9c move lemmas to Deflation.thy
Brian Huffman <brianh@cs.pdx.edu>
parents: 36452
diff changeset
   162
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   163
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   164
subsection \<open>Continuous embedding-projection pairs\<close>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   165
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   166
locale ep_pair =
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   167
  fixes e :: "'a \<rightarrow> 'b" and p :: "'b \<rightarrow> 'a"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   168
  assumes e_inverse [simp]: "\<And>x. p\<cdot>(e\<cdot>x) = x"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   169
  and e_p_below: "\<And>y. e\<cdot>(p\<cdot>y) \<sqsubseteq> y"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   170
begin
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   171
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   172
lemma e_below_iff [simp]: "e\<cdot>x \<sqsubseteq> e\<cdot>y \<longleftrightarrow> x \<sqsubseteq> y"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   173
proof
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   174
  assume "e\<cdot>x \<sqsubseteq> e\<cdot>y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   175
  then have "p\<cdot>(e\<cdot>x) \<sqsubseteq> p\<cdot>(e\<cdot>y)" by (rule monofun_cfun_arg)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   176
  then show "x \<sqsubseteq> y" by simp
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   177
next
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   178
  assume "x \<sqsubseteq> y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   179
  then show "e\<cdot>x \<sqsubseteq> e\<cdot>y" by (rule monofun_cfun_arg)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   180
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   181
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   182
lemma e_eq_iff [simp]: "e\<cdot>x = e\<cdot>y \<longleftrightarrow> x = y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   183
  unfolding po_eq_conv e_below_iff ..
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   184
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   185
lemma p_eq_iff: "e\<cdot>(p\<cdot>x) = x \<Longrightarrow> e\<cdot>(p\<cdot>y) = y \<Longrightarrow> p\<cdot>x = p\<cdot>y \<longleftrightarrow> x = y"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   186
  by (safe, erule subst, erule subst, simp)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   187
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   188
lemma p_inverse: "(\<exists>x. y = e\<cdot>x) \<longleftrightarrow> e\<cdot>(p\<cdot>y) = y"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   189
  by (auto, rule exI, erule sym)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   190
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   191
lemma e_below_iff_below_p: "e\<cdot>x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> p\<cdot>y"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   192
proof
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   193
  assume "e\<cdot>x \<sqsubseteq> y"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   194
  then have "p\<cdot>(e\<cdot>x) \<sqsubseteq> p\<cdot>y" by (rule monofun_cfun_arg)
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   195
  then show "x \<sqsubseteq> p\<cdot>y" by simp
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   196
next
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   197
  assume "x \<sqsubseteq> p\<cdot>y"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   198
  then have "e\<cdot>x \<sqsubseteq> e\<cdot>(p\<cdot>y)" by (rule monofun_cfun_arg)
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   199
  then show "e\<cdot>x \<sqsubseteq> y" using e_p_below by (rule below_trans)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   200
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   201
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   202
lemma compact_e_rev: "compact (e\<cdot>x) \<Longrightarrow> compact x"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   203
proof -
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   204
  assume "compact (e\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   205
  then have "adm (\<lambda>y. e\<cdot>x \<notsqsubseteq> y)" by (rule compactD)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   206
  then have "adm (\<lambda>y. e\<cdot>x \<notsqsubseteq> e\<cdot>y)" by (rule adm_subst [OF cont_Rep_cfun2])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   207
  then have "adm (\<lambda>y. x \<notsqsubseteq> y)" by simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   208
  then show "compact x" by (rule compactI)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   209
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   210
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   211
lemma compact_e:
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   212
  assumes "compact x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   213
  shows "compact (e\<cdot>x)"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   214
proof -
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   215
  from assms have "adm (\<lambda>y. x \<notsqsubseteq> y)" by (rule compactD)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   216
  then have "adm (\<lambda>y. x \<notsqsubseteq> p\<cdot>y)" by (rule adm_subst [OF cont_Rep_cfun2])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   217
  then have "adm (\<lambda>y. e\<cdot>x \<notsqsubseteq> y)" by (simp add: e_below_iff_below_p)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   218
  then show "compact (e\<cdot>x)" by (rule compactI)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   219
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   220
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   221
lemma compact_e_iff: "compact (e\<cdot>x) \<longleftrightarrow> compact x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   222
  by (rule iffI [OF compact_e_rev compact_e])
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   223
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   224
text \<open>Deflations from ep-pairs\<close>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   225
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   226
lemma deflation_e_p: "deflation (e oo p)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   227
  by (simp add: deflation.intro e_p_below)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   228
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   229
lemma deflation_e_d_p:
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
   230
  assumes "deflation d"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   231
  shows "deflation (e oo d oo p)"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   232
proof
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28613
diff changeset
   233
  interpret deflation d by fact
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   234
  fix x :: 'b
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   235
  show "(e oo d oo p)\<cdot>((e oo d oo p)\<cdot>x) = (e oo d oo p)\<cdot>x"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   236
    by (simp add: idem)
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   237
  show "(e oo d oo p)\<cdot>x \<sqsubseteq> x"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   238
    by (simp add: e_below_iff_below_p below)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   239
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   240
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   241
lemma finite_deflation_e_d_p:
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
   242
  assumes "finite_deflation d"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   243
  shows "finite_deflation (e oo d oo p)"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   244
proof
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28613
diff changeset
   245
  interpret finite_deflation d by fact
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   246
  fix x :: 'b
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   247
  show "(e oo d oo p)\<cdot>((e oo d oo p)\<cdot>x) = (e oo d oo p)\<cdot>x"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   248
    by (simp add: idem)
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   249
  show "(e oo d oo p)\<cdot>x \<sqsubseteq> x"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   250
    by (simp add: e_below_iff_below_p below)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   251
  have "finite ((\<lambda>x. e\<cdot>x) ` (\<lambda>x. d\<cdot>x) ` range (\<lambda>x. p\<cdot>x))"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   252
    by (simp add: finite_image)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   253
  then have "finite (range (\<lambda>x. (e oo d oo p)\<cdot>x))"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   254
    by (simp add: image_image)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   255
  then show "finite {x. (e oo d oo p)\<cdot>x = x}"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   256
    by (rule finite_range_imp_finite_fixes)
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   257
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   258
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   259
lemma deflation_p_d_e:
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
   260
  assumes "deflation d"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   261
  assumes d: "\<And>x. d\<cdot>x \<sqsubseteq> e\<cdot>(p\<cdot>x)"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   262
  shows "deflation (p oo d oo e)"
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
   263
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28613
diff changeset
   264
  interpret d: deflation d by fact
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   265
  have p_d_e_below: "(p oo d oo e)\<cdot>x \<sqsubseteq> x" for x
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   266
  proof -
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   267
    have "d\<cdot>(e\<cdot>x) \<sqsubseteq> e\<cdot>x"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   268
      by (rule d.below)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   269
    then have "p\<cdot>(d\<cdot>(e\<cdot>x)) \<sqsubseteq> p\<cdot>(e\<cdot>x)"
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   270
      by (rule monofun_cfun_arg)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   271
    then show ?thesis by simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   272
  qed
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
   273
  show ?thesis
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   274
  proof
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   275
    show "(p oo d oo e)\<cdot>x \<sqsubseteq> x" for x
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   276
      by (rule p_d_e_below)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   277
    show "(p oo d oo e)\<cdot>((p oo d oo e)\<cdot>x) = (p oo d oo e)\<cdot>x" for x
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   278
    proof (rule below_antisym)
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   279
      show "(p oo d oo e)\<cdot>((p oo d oo e)\<cdot>x) \<sqsubseteq> (p oo d oo e)\<cdot>x"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   280
        by (rule p_d_e_below)
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   281
      have "p\<cdot>(d\<cdot>(d\<cdot>(d\<cdot>(e\<cdot>x)))) \<sqsubseteq> p\<cdot>(d\<cdot>(e\<cdot>(p\<cdot>(d\<cdot>(e\<cdot>x)))))"
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   282
        by (intro monofun_cfun_arg d)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   283
      then have "p\<cdot>(d\<cdot>(e\<cdot>x)) \<sqsubseteq> p\<cdot>(d\<cdot>(e\<cdot>(p\<cdot>(d\<cdot>(e\<cdot>x)))))"
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   284
        by (simp only: d.idem)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   285
      then show "(p oo d oo e)\<cdot>x \<sqsubseteq> (p oo d oo e)\<cdot>((p oo d oo e)\<cdot>x)"
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   286
        by simp
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   287
    qed
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   288
  qed
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
   289
qed
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   290
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   291
lemma finite_deflation_p_d_e:
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
   292
  assumes "finite_deflation d"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   293
  assumes d: "\<And>x. d\<cdot>x \<sqsubseteq> e\<cdot>(p\<cdot>x)"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   294
  shows "finite_deflation (p oo d oo e)"
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
   295
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28613
diff changeset
   296
  interpret d: finite_deflation d by fact
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
   297
  show ?thesis
39973
c62b4ff97bfc add lemma finite_deflation_intro
Brian Huffman <brianh@cs.pdx.edu>
parents: 39971
diff changeset
   298
  proof (rule finite_deflation_intro)
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   299
    have "deflation d" ..
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   300
    then show "deflation (p oo d oo e)"
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   301
      using d by (rule deflation_p_d_e)
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   302
  next
39973
c62b4ff97bfc add lemma finite_deflation_intro
Brian Huffman <brianh@cs.pdx.edu>
parents: 39971
diff changeset
   303
    have "finite ((\<lambda>x. d\<cdot>x) ` range (\<lambda>x. e\<cdot>x))"
c62b4ff97bfc add lemma finite_deflation_intro
Brian Huffman <brianh@cs.pdx.edu>
parents: 39971
diff changeset
   304
      by (rule d.finite_image)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   305
    then have "finite ((\<lambda>x. p\<cdot>x) ` (\<lambda>x. d\<cdot>x) ` range (\<lambda>x. e\<cdot>x))"
39973
c62b4ff97bfc add lemma finite_deflation_intro
Brian Huffman <brianh@cs.pdx.edu>
parents: 39971
diff changeset
   306
      by (rule finite_imageI)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   307
    then have "finite (range (\<lambda>x. (p oo d oo e)\<cdot>x))"
39973
c62b4ff97bfc add lemma finite_deflation_intro
Brian Huffman <brianh@cs.pdx.edu>
parents: 39971
diff changeset
   308
      by (simp add: image_image)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   309
    then show "finite {x. (p oo d oo e)\<cdot>x = x}"
39973
c62b4ff97bfc add lemma finite_deflation_intro
Brian Huffman <brianh@cs.pdx.edu>
parents: 39971
diff changeset
   310
      by (rule finite_range_imp_finite_fixes)
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   311
  qed
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27681
diff changeset
   312
qed
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   313
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   314
end
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   315
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   316
subsection \<open>Uniqueness of ep-pairs\<close>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   317
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   318
lemma ep_pair_unique_e_lemma:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   319
  assumes 1: "ep_pair e1 p"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   320
    and 2: "ep_pair e2 p"
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   321
  shows "e1 \<sqsubseteq> e2"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 39985
diff changeset
   322
proof (rule cfun_belowI)
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   323
  fix x
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   324
  have "e1\<cdot>(p\<cdot>(e2\<cdot>x)) \<sqsubseteq> e2\<cdot>x"
35168
07b3112e464b fix warnings about duplicate simp rules
huffman
parents: 33503
diff changeset
   325
    by (rule ep_pair.e_p_below [OF 1])
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   326
  then show "e1\<cdot>x \<sqsubseteq> e2\<cdot>x"
35168
07b3112e464b fix warnings about duplicate simp rules
huffman
parents: 33503
diff changeset
   327
    by (simp only: ep_pair.e_inverse [OF 2])
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   328
qed
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   329
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   330
lemma ep_pair_unique_e: "ep_pair e1 p \<Longrightarrow> ep_pair e2 p \<Longrightarrow> e1 = e2"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   331
  by (fast intro: below_antisym elim: ep_pair_unique_e_lemma)
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   332
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   333
lemma ep_pair_unique_p_lemma:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   334
  assumes 1: "ep_pair e p1"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   335
    and 2: "ep_pair e p2"
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   336
  shows "p1 \<sqsubseteq> p2"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 39985
diff changeset
   337
proof (rule cfun_belowI)
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   338
  fix x
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   339
  have "e\<cdot>(p1\<cdot>x) \<sqsubseteq> x"
35168
07b3112e464b fix warnings about duplicate simp rules
huffman
parents: 33503
diff changeset
   340
    by (rule ep_pair.e_p_below [OF 1])
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   341
  then have "p2\<cdot>(e\<cdot>(p1\<cdot>x)) \<sqsubseteq> p2\<cdot>x"
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   342
    by (rule monofun_cfun_arg)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   343
  then show "p1\<cdot>x \<sqsubseteq> p2\<cdot>x"
35168
07b3112e464b fix warnings about duplicate simp rules
huffman
parents: 33503
diff changeset
   344
    by (simp only: ep_pair.e_inverse [OF 2])
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   345
qed
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   346
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   347
lemma ep_pair_unique_p: "ep_pair e p1 \<Longrightarrow> ep_pair e p2 \<Longrightarrow> p1 = p2"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   348
  by (fast intro: below_antisym elim: ep_pair_unique_p_lemma)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   349
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   350
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   351
subsection \<open>Composing ep-pairs\<close>
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   352
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   353
lemma ep_pair_ID_ID: "ep_pair ID ID"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   354
  by standard simp_all
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   355
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   356
lemma ep_pair_comp:
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   357
  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   358
  shows "ep_pair (e2 oo e1) (p1 oo p2)"
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   359
proof
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28613
diff changeset
   360
  interpret ep1: ep_pair e1 p1 by fact
e90d9d51106b More porting to new locales.
ballarin
parents: 28613
diff changeset
   361
  interpret ep2: ep_pair e2 p2 by fact
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   362
  fix x y
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   363
  show "(p1 oo p2)\<cdot>((e2 oo e1)\<cdot>x) = x"
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   364
    by simp
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   365
  have "e1\<cdot>(p1\<cdot>(p2\<cdot>y)) \<sqsubseteq> p2\<cdot>y"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   366
    by (rule ep1.e_p_below)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   367
  then have "e2\<cdot>(e1\<cdot>(p1\<cdot>(p2\<cdot>y))) \<sqsubseteq> e2\<cdot>(p2\<cdot>y)"
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   368
    by (rule monofun_cfun_arg)
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   369
  also have "e2\<cdot>(p2\<cdot>y) \<sqsubseteq> y"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   370
    by (rule ep2.e_p_below)
28613
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   371
  finally show "(e2 oo e1)\<cdot>((p1 oo p2)\<cdot>y) \<sqsubseteq> y"
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   372
    by simp
15a41d3fa959 rewrite more proofs in Isar style
huffman
parents: 28611
diff changeset
   373
qed
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   374
46868
6c250adbe101 eliminated old-fashioned 'constrains' element;
wenzelm
parents: 42151
diff changeset
   375
locale pcpo_ep_pair = ep_pair e p
6c250adbe101 eliminated old-fashioned 'constrains' element;
wenzelm
parents: 42151
diff changeset
   376
  for e :: "'a::pcpo \<rightarrow> 'b::pcpo"
6c250adbe101 eliminated old-fashioned 'constrains' element;
wenzelm
parents: 42151
diff changeset
   377
  and p :: "'b::pcpo \<rightarrow> 'a::pcpo"
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   378
begin
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   379
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   380
lemma e_strict [simp]: "e\<cdot>\<bottom> = \<bottom>"
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   381
proof -
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   382
  have "\<bottom> \<sqsubseteq> p\<cdot>\<bottom>" by (rule minimal)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   383
  then have "e\<cdot>\<bottom> \<sqsubseteq> e\<cdot>(p\<cdot>\<bottom>)" by (rule monofun_cfun_arg)
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29252
diff changeset
   384
  also have "e\<cdot>(p\<cdot>\<bottom>) \<sqsubseteq> \<bottom>" by (rule e_p_below)
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   385
  finally show "e\<cdot>\<bottom> = \<bottom>" by simp
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   386
qed
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   387
40321
d065b195ec89 rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
huffman
parents: 40002
diff changeset
   388
lemma e_bottom_iff [simp]: "e\<cdot>x = \<bottom> \<longleftrightarrow> x = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   389
  by (rule e_eq_iff [where y="\<bottom>", unfolded e_strict])
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   390
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   391
lemma e_defined: "x \<noteq> \<bottom> \<Longrightarrow> e\<cdot>x \<noteq> \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   392
  by simp
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   393
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   394
lemma p_strict [simp]: "p\<cdot>\<bottom> = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 65380
diff changeset
   395
  by (rule e_inverse [where x="\<bottom>", unfolded e_strict])
27401
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   396
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   397
lemmas stricts = e_strict p_strict
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   398
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   399
end
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   400
4edc81f93e35 New theory of deflations and embedding-projection pairs
huffman
parents:
diff changeset
   401
end