src/HOLCF/ConvexPD.thy
author ballarin
Fri, 19 Dec 2008 11:57:21 +0100
changeset 29244 95d591908d8d
parent 29237 e90d9d51106b
child 29252 ea97aa6aeba2
permissions -rw-r--r--
More porting to new locales.
huffman@25904
     1
(*  Title:      HOLCF/ConvexPD.thy
huffman@25904
     2
    Author:     Brian Huffman
huffman@25904
     3
*)
huffman@25904
     4
huffman@25904
     5
header {* Convex powerdomain *}
huffman@25904
     6
huffman@25904
     7
theory ConvexPD
huffman@25904
     8
imports UpperPD LowerPD
huffman@25904
     9
begin
huffman@25904
    10
huffman@25904
    11
subsection {* Basis preorder *}
huffman@25904
    12
huffman@25904
    13
definition
huffman@25904
    14
  convex_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix "\<le>\<natural>" 50) where
huffman@25904
    15
  "convex_le = (\<lambda>u v. u \<le>\<sharp> v \<and> u \<le>\<flat> v)"
huffman@25904
    16
huffman@25904
    17
lemma convex_le_refl [simp]: "t \<le>\<natural> t"
huffman@25904
    18
unfolding convex_le_def by (fast intro: upper_le_refl lower_le_refl)
huffman@25904
    19
huffman@25904
    20
lemma convex_le_trans: "\<lbrakk>t \<le>\<natural> u; u \<le>\<natural> v\<rbrakk> \<Longrightarrow> t \<le>\<natural> v"
huffman@25904
    21
unfolding convex_le_def by (fast intro: upper_le_trans lower_le_trans)
huffman@25904
    22
ballarin@29237
    23
interpretation convex_le!: preorder convex_le
huffman@25904
    24
by (rule preorder.intro, rule convex_le_refl, rule convex_le_trans)
huffman@25904
    25
huffman@25904
    26
lemma upper_le_minimal [simp]: "PDUnit compact_bot \<le>\<natural> t"
huffman@25904
    27
unfolding convex_le_def Rep_PDUnit by simp
huffman@25904
    28
huffman@26420
    29
lemma PDUnit_convex_mono: "x \<sqsubseteq> y \<Longrightarrow> PDUnit x \<le>\<natural> PDUnit y"
huffman@25904
    30
unfolding convex_le_def by (fast intro: PDUnit_upper_mono PDUnit_lower_mono)
huffman@25904
    31
huffman@25904
    32
lemma PDPlus_convex_mono: "\<lbrakk>s \<le>\<natural> t; u \<le>\<natural> v\<rbrakk> \<Longrightarrow> PDPlus s u \<le>\<natural> PDPlus t v"
huffman@25904
    33
unfolding convex_le_def by (fast intro: PDPlus_upper_mono PDPlus_lower_mono)
huffman@25904
    34
huffman@25904
    35
lemma convex_le_PDUnit_PDUnit_iff [simp]:
huffman@26420
    36
  "(PDUnit a \<le>\<natural> PDUnit b) = a \<sqsubseteq> b"
huffman@25904
    37
unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit by fast
huffman@25904
    38
huffman@25904
    39
lemma convex_le_PDUnit_lemma1:
huffman@26420
    40
  "(PDUnit a \<le>\<natural> t) = (\<forall>b\<in>Rep_pd_basis t. a \<sqsubseteq> b)"
huffman@25904
    41
unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit
huffman@25904
    42
using Rep_pd_basis_nonempty [of t, folded ex_in_conv] by fast
huffman@25904
    43
huffman@25904
    44
lemma convex_le_PDUnit_PDPlus_iff [simp]:
huffman@25904
    45
  "(PDUnit a \<le>\<natural> PDPlus t u) = (PDUnit a \<le>\<natural> t \<and> PDUnit a \<le>\<natural> u)"
huffman@25904
    46
unfolding convex_le_PDUnit_lemma1 Rep_PDPlus by fast
huffman@25904
    47
huffman@25904
    48
lemma convex_le_PDUnit_lemma2:
huffman@26420
    49
  "(t \<le>\<natural> PDUnit b) = (\<forall>a\<in>Rep_pd_basis t. a \<sqsubseteq> b)"
huffman@25904
    50
unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit
huffman@25904
    51
using Rep_pd_basis_nonempty [of t, folded ex_in_conv] by fast
huffman@25904
    52
huffman@25904
    53
lemma convex_le_PDPlus_PDUnit_iff [simp]:
huffman@25904
    54
  "(PDPlus t u \<le>\<natural> PDUnit a) = (t \<le>\<natural> PDUnit a \<and> u \<le>\<natural> PDUnit a)"
huffman@25904
    55
unfolding convex_le_PDUnit_lemma2 Rep_PDPlus by fast
huffman@25904
    56
huffman@25904
    57
lemma convex_le_PDPlus_lemma:
huffman@25904
    58
  assumes z: "PDPlus t u \<le>\<natural> z"
huffman@25904
    59
  shows "\<exists>v w. z = PDPlus v w \<and> t \<le>\<natural> v \<and> u \<le>\<natural> w"
huffman@25904
    60
proof (intro exI conjI)
huffman@26420
    61
  let ?A = "{b\<in>Rep_pd_basis z. \<exists>a\<in>Rep_pd_basis t. a \<sqsubseteq> b}"
huffman@26420
    62
  let ?B = "{b\<in>Rep_pd_basis z. \<exists>a\<in>Rep_pd_basis u. a \<sqsubseteq> b}"
huffman@25904
    63
  let ?v = "Abs_pd_basis ?A"
huffman@25904
    64
  let ?w = "Abs_pd_basis ?B"
huffman@25904
    65
  have Rep_v: "Rep_pd_basis ?v = ?A"
huffman@25904
    66
    apply (rule Abs_pd_basis_inverse)
huffman@25904
    67
    apply (rule Rep_pd_basis_nonempty [of t, folded ex_in_conv, THEN exE])
huffman@25904
    68
    apply (cut_tac z, simp only: convex_le_def lower_le_def, clarify)
huffman@25904
    69
    apply (drule_tac x=x in bspec, simp add: Rep_PDPlus, erule bexE)
huffman@25904
    70
    apply (simp add: pd_basis_def)
huffman@25904
    71
    apply fast
huffman@25904
    72
    done
huffman@25904
    73
  have Rep_w: "Rep_pd_basis ?w = ?B"
huffman@25904
    74
    apply (rule Abs_pd_basis_inverse)
huffman@25904
    75
    apply (rule Rep_pd_basis_nonempty [of u, folded ex_in_conv, THEN exE])
huffman@25904
    76
    apply (cut_tac z, simp only: convex_le_def lower_le_def, clarify)
huffman@25904
    77
    apply (drule_tac x=x in bspec, simp add: Rep_PDPlus, erule bexE)
huffman@25904
    78
    apply (simp add: pd_basis_def)
huffman@25904
    79
    apply fast
huffman@25904
    80
    done
huffman@25904
    81
  show "z = PDPlus ?v ?w"
huffman@25904
    82
    apply (insert z)
huffman@25904
    83
    apply (simp add: convex_le_def, erule conjE)
huffman@25904
    84
    apply (simp add: Rep_pd_basis_inject [symmetric] Rep_PDPlus)
huffman@25904
    85
    apply (simp add: Rep_v Rep_w)
huffman@25904
    86
    apply (rule equalityI)
huffman@25904
    87
     apply (rule subsetI)
huffman@25904
    88
     apply (simp only: upper_le_def)
huffman@25904
    89
     apply (drule (1) bspec, erule bexE)
huffman@25904
    90
     apply (simp add: Rep_PDPlus)
huffman@25904
    91
     apply fast
huffman@25904
    92
    apply fast
huffman@25904
    93
    done
huffman@25904
    94
  show "t \<le>\<natural> ?v" "u \<le>\<natural> ?w"
huffman@25904
    95
   apply (insert z)
huffman@25904
    96
   apply (simp_all add: convex_le_def upper_le_def lower_le_def Rep_PDPlus Rep_v Rep_w)
huffman@25904
    97
   apply fast+
huffman@25904
    98
   done
huffman@25904
    99
qed
huffman@25904
   100
huffman@25904
   101
lemma convex_le_induct [induct set: convex_le]:
huffman@25904
   102
  assumes le: "t \<le>\<natural> u"
huffman@25904
   103
  assumes 2: "\<And>t u v. \<lbrakk>P t u; P u v\<rbrakk> \<Longrightarrow> P t v"
huffman@26420
   104
  assumes 3: "\<And>a b. a \<sqsubseteq> b \<Longrightarrow> P (PDUnit a) (PDUnit b)"
huffman@25904
   105
  assumes 4: "\<And>t u v w. \<lbrakk>P t v; P u w\<rbrakk> \<Longrightarrow> P (PDPlus t u) (PDPlus v w)"
huffman@25904
   106
  shows "P t u"
huffman@25904
   107
using le apply (induct t arbitrary: u rule: pd_basis_induct)
huffman@25904
   108
apply (erule rev_mp)
huffman@25904
   109
apply (induct_tac u rule: pd_basis_induct1)
huffman@25904
   110
apply (simp add: 3)
huffman@25904
   111
apply (simp, clarify, rename_tac a b t)
huffman@25904
   112
apply (subgoal_tac "P (PDPlus (PDUnit a) (PDUnit a)) (PDPlus (PDUnit b) t)")
huffman@25904
   113
apply (simp add: PDPlus_absorb)
huffman@25904
   114
apply (erule (1) 4 [OF 3])
huffman@25904
   115
apply (drule convex_le_PDPlus_lemma, clarify)
huffman@25904
   116
apply (simp add: 4)
huffman@25904
   117
done
huffman@25904
   118
huffman@27405
   119
lemma pd_take_convex_chain:
huffman@27405
   120
  "pd_take n t \<le>\<natural> pd_take (Suc n) t"
huffman@25904
   121
apply (induct t rule: pd_basis_induct)
huffman@27289
   122
apply (simp add: compact_basis.take_chain)
huffman@25904
   123
apply (simp add: PDPlus_convex_mono)
huffman@25904
   124
done
huffman@25904
   125
huffman@27405
   126
lemma pd_take_convex_le: "pd_take i t \<le>\<natural> t"
huffman@25904
   127
apply (induct t rule: pd_basis_induct)
huffman@27289
   128
apply (simp add: compact_basis.take_less)
huffman@25904
   129
apply (simp add: PDPlus_convex_mono)
huffman@25904
   130
done
huffman@25904
   131
huffman@27405
   132
lemma pd_take_convex_mono:
huffman@27405
   133
  "t \<le>\<natural> u \<Longrightarrow> pd_take n t \<le>\<natural> pd_take n u"
huffman@25904
   134
apply (erule convex_le_induct)
huffman@25904
   135
apply (erule (1) convex_le_trans)
huffman@27289
   136
apply (simp add: compact_basis.take_mono)
huffman@25904
   137
apply (simp add: PDPlus_convex_mono)
huffman@25904
   138
done
huffman@25904
   139
huffman@25904
   140
huffman@25904
   141
subsection {* Type definition *}
huffman@25904
   142
huffman@27373
   143
typedef (open) 'a convex_pd =
huffman@27373
   144
  "{S::'a pd_basis set. convex_le.ideal S}"
huffman@27373
   145
by (fast intro: convex_le.ideal_principal)
huffman@25904
   146
huffman@27373
   147
instantiation convex_pd :: (profinite) sq_ord
huffman@27373
   148
begin
huffman@27373
   149
huffman@27373
   150
definition
huffman@27373
   151
  "x \<sqsubseteq> y \<longleftrightarrow> Rep_convex_pd x \<subseteq> Rep_convex_pd y"
huffman@27373
   152
huffman@27373
   153
instance ..
huffman@27373
   154
end
huffman@27373
   155
huffman@27373
   156
instance convex_pd :: (profinite) po
huffman@27373
   157
by (rule convex_le.typedef_ideal_po
huffman@27373
   158
    [OF type_definition_convex_pd sq_le_convex_pd_def])
huffman@27373
   159
huffman@27373
   160
instance convex_pd :: (profinite) cpo
huffman@27373
   161
by (rule convex_le.typedef_ideal_cpo
huffman@27373
   162
    [OF type_definition_convex_pd sq_le_convex_pd_def])
huffman@27373
   163
huffman@27373
   164
lemma Rep_convex_pd_lub:
huffman@27373
   165
  "chain Y \<Longrightarrow> Rep_convex_pd (\<Squnion>i. Y i) = (\<Union>i. Rep_convex_pd (Y i))"
huffman@27373
   166
by (rule convex_le.typedef_ideal_rep_contlub
huffman@27373
   167
    [OF type_definition_convex_pd sq_le_convex_pd_def])
huffman@27373
   168
huffman@27373
   169
lemma ideal_Rep_convex_pd: "convex_le.ideal (Rep_convex_pd xs)"
huffman@26927
   170
by (rule Rep_convex_pd [unfolded mem_Collect_eq])
huffman@25904
   171
huffman@25904
   172
definition
huffman@25904
   173
  convex_principal :: "'a pd_basis \<Rightarrow> 'a convex_pd" where
huffman@27373
   174
  "convex_principal t = Abs_convex_pd {u. u \<le>\<natural> t}"
huffman@25904
   175
huffman@25904
   176
lemma Rep_convex_principal:
huffman@27373
   177
  "Rep_convex_pd (convex_principal t) = {u. u \<le>\<natural> t}"
huffman@25904
   178
unfolding convex_principal_def
huffman@27297
   179
by (simp add: Abs_convex_pd_inverse convex_le.ideal_principal)
huffman@25904
   180
ballarin@29237
   181
interpretation convex_pd!:
ballarin@29237
   182
  ideal_completion convex_le pd_take convex_principal Rep_convex_pd
huffman@25904
   183
apply unfold_locales
huffman@27405
   184
apply (rule pd_take_convex_le)
huffman@27405
   185
apply (rule pd_take_idem)
huffman@27405
   186
apply (erule pd_take_convex_mono)
huffman@27405
   187
apply (rule pd_take_convex_chain)
huffman@27405
   188
apply (rule finite_range_pd_take)
huffman@27405
   189
apply (rule pd_take_covers)
huffman@26420
   190
apply (rule ideal_Rep_convex_pd)
huffman@27373
   191
apply (erule Rep_convex_pd_lub)
huffman@26420
   192
apply (rule Rep_convex_principal)
huffman@27373
   193
apply (simp only: sq_le_convex_pd_def)
huffman@25904
   194
done
huffman@25904
   195
huffman@27289
   196
text {* Convex powerdomain is pointed *}
huffman@25904
   197
huffman@25904
   198
lemma convex_pd_minimal: "convex_principal (PDUnit compact_bot) \<sqsubseteq> ys"
huffman@25904
   199
by (induct ys rule: convex_pd.principal_induct, simp, simp)
huffman@25904
   200
huffman@25904
   201
instance convex_pd :: (bifinite) pcpo
huffman@26927
   202
by intro_classes (fast intro: convex_pd_minimal)
huffman@25904
   203
huffman@25904
   204
lemma inst_convex_pd_pcpo: "\<bottom> = convex_principal (PDUnit compact_bot)"
huffman@25904
   205
by (rule convex_pd_minimal [THEN UU_I, symmetric])
huffman@25904
   206
huffman@27289
   207
text {* Convex powerdomain is profinite *}
huffman@25904
   208
huffman@26962
   209
instantiation convex_pd :: (profinite) profinite
huffman@26962
   210
begin
huffman@25904
   211
huffman@26962
   212
definition
huffman@26962
   213
  approx_convex_pd_def: "approx = convex_pd.completion_approx"
huffman@26927
   214
huffman@26962
   215
instance
huffman@26927
   216
apply (intro_classes, unfold approx_convex_pd_def)
huffman@27310
   217
apply (rule convex_pd.chain_completion_approx)
huffman@26927
   218
apply (rule convex_pd.lub_completion_approx)
huffman@26927
   219
apply (rule convex_pd.completion_approx_idem)
huffman@26927
   220
apply (rule convex_pd.finite_fixes_completion_approx)
huffman@26927
   221
done
huffman@26927
   222
huffman@26962
   223
end
huffman@26962
   224
huffman@26927
   225
instance convex_pd :: (bifinite) bifinite ..
huffman@25904
   226
huffman@25904
   227
lemma approx_convex_principal [simp]:
huffman@27405
   228
  "approx n\<cdot>(convex_principal t) = convex_principal (pd_take n t)"
huffman@25904
   229
unfolding approx_convex_pd_def
huffman@26927
   230
by (rule convex_pd.completion_approx_principal)
huffman@25904
   231
huffman@25904
   232
lemma approx_eq_convex_principal:
huffman@27405
   233
  "\<exists>t\<in>Rep_convex_pd xs. approx n\<cdot>xs = convex_principal (pd_take n t)"
huffman@25904
   234
unfolding approx_convex_pd_def
huffman@26927
   235
by (rule convex_pd.completion_approx_eq_principal)
huffman@26407
   236
huffman@25904
   237
huffman@26927
   238
subsection {* Monadic unit and plus *}
huffman@25904
   239
huffman@25904
   240
definition
huffman@25904
   241
  convex_unit :: "'a \<rightarrow> 'a convex_pd" where
huffman@25904
   242
  "convex_unit = compact_basis.basis_fun (\<lambda>a. convex_principal (PDUnit a))"
huffman@25904
   243
huffman@25904
   244
definition
huffman@25904
   245
  convex_plus :: "'a convex_pd \<rightarrow> 'a convex_pd \<rightarrow> 'a convex_pd" where
huffman@25904
   246
  "convex_plus = convex_pd.basis_fun (\<lambda>t. convex_pd.basis_fun (\<lambda>u.
huffman@25904
   247
      convex_principal (PDPlus t u)))"
huffman@25904
   248
huffman@25904
   249
abbreviation
huffman@25904
   250
  convex_add :: "'a convex_pd \<Rightarrow> 'a convex_pd \<Rightarrow> 'a convex_pd"
huffman@25904
   251
    (infixl "+\<natural>" 65) where
huffman@25904
   252
  "xs +\<natural> ys == convex_plus\<cdot>xs\<cdot>ys"
huffman@25904
   253
huffman@26927
   254
syntax
huffman@26927
   255
  "_convex_pd" :: "args \<Rightarrow> 'a convex_pd" ("{_}\<natural>")
huffman@26927
   256
huffman@26927
   257
translations
huffman@26927
   258
  "{x,xs}\<natural>" == "{x}\<natural> +\<natural> {xs}\<natural>"
huffman@26927
   259
  "{x}\<natural>" == "CONST convex_unit\<cdot>x"
huffman@26927
   260
huffman@26927
   261
lemma convex_unit_Rep_compact_basis [simp]:
huffman@26927
   262
  "{Rep_compact_basis a}\<natural> = convex_principal (PDUnit a)"
huffman@26927
   263
unfolding convex_unit_def
huffman@27289
   264
by (simp add: compact_basis.basis_fun_principal PDUnit_convex_mono)
huffman@26927
   265
huffman@25904
   266
lemma convex_plus_principal [simp]:
huffman@26927
   267
  "convex_principal t +\<natural> convex_principal u = convex_principal (PDPlus t u)"
huffman@25904
   268
unfolding convex_plus_def
huffman@25904
   269
by (simp add: convex_pd.basis_fun_principal
huffman@25904
   270
    convex_pd.basis_fun_mono PDPlus_convex_mono)
huffman@25904
   271
huffman@26927
   272
lemma approx_convex_unit [simp]:
huffman@26927
   273
  "approx n\<cdot>{x}\<natural> = {approx n\<cdot>x}\<natural>"
huffman@27289
   274
apply (induct x rule: compact_basis.principal_induct, simp)
huffman@26927
   275
apply (simp add: approx_Rep_compact_basis)
huffman@26927
   276
done
huffman@26927
   277
huffman@25904
   278
lemma approx_convex_plus [simp]:
huffman@26927
   279
  "approx n\<cdot>(xs +\<natural> ys) = approx n\<cdot>xs +\<natural> approx n\<cdot>ys"
huffman@27289
   280
by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp)
huffman@25904
   281
huffman@25904
   282
lemma convex_plus_assoc:
huffman@26927
   283
  "(xs +\<natural> ys) +\<natural> zs = xs +\<natural> (ys +\<natural> zs)"
huffman@27289
   284
apply (induct xs ys arbitrary: zs rule: convex_pd.principal_induct2, simp, simp)
huffman@27289
   285
apply (rule_tac x=zs in convex_pd.principal_induct, simp)
huffman@25904
   286
apply (simp add: PDPlus_assoc)
huffman@25904
   287
done
huffman@25904
   288
huffman@26927
   289
lemma convex_plus_commute: "xs +\<natural> ys = ys +\<natural> xs"
huffman@27289
   290
apply (induct xs ys rule: convex_pd.principal_induct2, simp, simp)
huffman@26927
   291
apply (simp add: PDPlus_commute)
huffman@26927
   292
done
huffman@26927
   293
huffman@26927
   294
lemma convex_plus_absorb: "xs +\<natural> xs = xs"
huffman@27289
   295
apply (induct xs rule: convex_pd.principal_induct, simp)
huffman@25904
   296
apply (simp add: PDPlus_absorb)
huffman@25904
   297
done
huffman@25904
   298
ballarin@29244
   299
class_interpretation aci_convex_plus: ab_semigroup_idem_mult ["op +\<natural>"]
huffman@26927
   300
  by unfold_locales
huffman@26927
   301
    (rule convex_plus_assoc convex_plus_commute convex_plus_absorb)+
huffman@26927
   302
huffman@26927
   303
lemma convex_plus_left_commute: "xs +\<natural> (ys +\<natural> zs) = ys +\<natural> (xs +\<natural> zs)"
huffman@26927
   304
by (rule aci_convex_plus.mult_left_commute)
huffman@26927
   305
huffman@26927
   306
lemma convex_plus_left_absorb: "xs +\<natural> (xs +\<natural> ys) = xs +\<natural> ys"
huffman@26927
   307
by (rule aci_convex_plus.mult_left_idem)
huffman@26927
   308
huffman@26927
   309
lemmas convex_plus_aci = aci_convex_plus.mult_ac_idem
huffman@26927
   310
huffman@25904
   311
lemma convex_unit_less_plus_iff [simp]:
huffman@26927
   312
  "{x}\<natural> \<sqsubseteq> ys +\<natural> zs \<longleftrightarrow> {x}\<natural> \<sqsubseteq> ys \<and> {x}\<natural> \<sqsubseteq> zs"
huffman@25904
   313
 apply (rule iffI)
huffman@25904
   314
  apply (subgoal_tac
huffman@26927
   315
    "adm (\<lambda>f. f\<cdot>{x}\<natural> \<sqsubseteq> f\<cdot>ys \<and> f\<cdot>{x}\<natural> \<sqsubseteq> f\<cdot>zs)")
huffman@25925
   316
   apply (drule admD, rule chain_approx)
huffman@25904
   317
    apply (drule_tac f="approx i" in monofun_cfun_arg)
huffman@27289
   318
    apply (cut_tac x="approx i\<cdot>x" in compact_basis.compact_imp_principal, simp)
huffman@27289
   319
    apply (cut_tac x="approx i\<cdot>ys" in convex_pd.compact_imp_principal, simp)
huffman@27289
   320
    apply (cut_tac x="approx i\<cdot>zs" in convex_pd.compact_imp_principal, simp)
huffman@25904
   321
    apply (clarify, simp)
huffman@25904
   322
   apply simp
huffman@25904
   323
  apply simp
huffman@25904
   324
 apply (erule conjE)
huffman@26927
   325
 apply (subst convex_plus_absorb [of "{x}\<natural>", symmetric])
huffman@25904
   326
 apply (erule (1) monofun_cfun [OF monofun_cfun_arg])
huffman@25904
   327
done
huffman@25904
   328
huffman@25904
   329
lemma convex_plus_less_unit_iff [simp]:
huffman@26927
   330
  "xs +\<natural> ys \<sqsubseteq> {z}\<natural> \<longleftrightarrow> xs \<sqsubseteq> {z}\<natural> \<and> ys \<sqsubseteq> {z}\<natural>"
huffman@25904
   331
 apply (rule iffI)
huffman@25904
   332
  apply (subgoal_tac
huffman@26927
   333
    "adm (\<lambda>f. f\<cdot>xs \<sqsubseteq> f\<cdot>{z}\<natural> \<and> f\<cdot>ys \<sqsubseteq> f\<cdot>{z}\<natural>)")
huffman@25925
   334
   apply (drule admD, rule chain_approx)
huffman@25904
   335
    apply (drule_tac f="approx i" in monofun_cfun_arg)
huffman@27289
   336
    apply (cut_tac x="approx i\<cdot>xs" in convex_pd.compact_imp_principal, simp)
huffman@27289
   337
    apply (cut_tac x="approx i\<cdot>ys" in convex_pd.compact_imp_principal, simp)
huffman@27289
   338
    apply (cut_tac x="approx i\<cdot>z" in compact_basis.compact_imp_principal, simp)
huffman@25904
   339
    apply (clarify, simp)
huffman@25904
   340
   apply simp
huffman@25904
   341
  apply simp
huffman@25904
   342
 apply (erule conjE)
huffman@26927
   343
 apply (subst convex_plus_absorb [of "{z}\<natural>", symmetric])
huffman@25904
   344
 apply (erule (1) monofun_cfun [OF monofun_cfun_arg])
huffman@25904
   345
done
huffman@25904
   346
huffman@26927
   347
lemma convex_unit_less_iff [simp]: "{x}\<natural> \<sqsubseteq> {y}\<natural> \<longleftrightarrow> x \<sqsubseteq> y"
huffman@26927
   348
 apply (rule iffI)
huffman@27309
   349
  apply (rule profinite_less_ext)
huffman@26927
   350
  apply (drule_tac f="approx i" in monofun_cfun_arg, simp)
huffman@27289
   351
  apply (cut_tac x="approx i\<cdot>x" in compact_basis.compact_imp_principal, simp)
huffman@27289
   352
  apply (cut_tac x="approx i\<cdot>y" in compact_basis.compact_imp_principal, simp)
huffman@27289
   353
  apply clarsimp
huffman@26927
   354
 apply (erule monofun_cfun_arg)
huffman@26927
   355
done
huffman@26927
   356
huffman@26927
   357
lemma convex_unit_eq_iff [simp]: "{x}\<natural> = {y}\<natural> \<longleftrightarrow> x = y"
huffman@26927
   358
unfolding po_eq_conv by simp
huffman@26927
   359
huffman@26927
   360
lemma convex_unit_strict [simp]: "{\<bottom>}\<natural> = \<bottom>"
huffman@26927
   361
unfolding inst_convex_pd_pcpo Rep_compact_bot [symmetric] by simp
huffman@26927
   362
huffman@26927
   363
lemma convex_unit_strict_iff [simp]: "{x}\<natural> = \<bottom> \<longleftrightarrow> x = \<bottom>"
huffman@26927
   364
unfolding convex_unit_strict [symmetric] by (rule convex_unit_eq_iff)
huffman@26927
   365
huffman@26927
   366
lemma compact_convex_unit_iff [simp]:
huffman@26927
   367
  "compact {x}\<natural> \<longleftrightarrow> compact x"
huffman@27309
   368
unfolding profinite_compact_iff by simp
huffman@26927
   369
huffman@26927
   370
lemma compact_convex_plus [simp]:
huffman@26927
   371
  "\<lbrakk>compact xs; compact ys\<rbrakk> \<Longrightarrow> compact (xs +\<natural> ys)"
huffman@27289
   372
by (auto dest!: convex_pd.compact_imp_principal)
huffman@26927
   373
huffman@25904
   374
huffman@25904
   375
subsection {* Induction rules *}
huffman@25904
   376
huffman@25904
   377
lemma convex_pd_induct1:
huffman@25904
   378
  assumes P: "adm P"
huffman@26927
   379
  assumes unit: "\<And>x. P {x}\<natural>"
huffman@26927
   380
  assumes insert: "\<And>x ys. \<lbrakk>P {x}\<natural>; P ys\<rbrakk> \<Longrightarrow> P ({x}\<natural> +\<natural> ys)"
huffman@25904
   381
  shows "P (xs::'a convex_pd)"
huffman@27289
   382
apply (induct xs rule: convex_pd.principal_induct, rule P)
huffman@27289
   383
apply (induct_tac a rule: pd_basis_induct1)
huffman@25904
   384
apply (simp only: convex_unit_Rep_compact_basis [symmetric])
huffman@25904
   385
apply (rule unit)
huffman@25904
   386
apply (simp only: convex_unit_Rep_compact_basis [symmetric]
huffman@25904
   387
                  convex_plus_principal [symmetric])
huffman@25904
   388
apply (erule insert [OF unit])
huffman@25904
   389
done
huffman@25904
   390
huffman@25904
   391
lemma convex_pd_induct:
huffman@25904
   392
  assumes P: "adm P"
huffman@26927
   393
  assumes unit: "\<And>x. P {x}\<natural>"
huffman@26927
   394
  assumes plus: "\<And>xs ys. \<lbrakk>P xs; P ys\<rbrakk> \<Longrightarrow> P (xs +\<natural> ys)"
huffman@25904
   395
  shows "P (xs::'a convex_pd)"
huffman@27289
   396
apply (induct xs rule: convex_pd.principal_induct, rule P)
huffman@27289
   397
apply (induct_tac a rule: pd_basis_induct)
huffman@25904
   398
apply (simp only: convex_unit_Rep_compact_basis [symmetric] unit)
huffman@25904
   399
apply (simp only: convex_plus_principal [symmetric] plus)
huffman@25904
   400
done
huffman@25904
   401
huffman@25904
   402
huffman@25904
   403
subsection {* Monadic bind *}
huffman@25904
   404
huffman@25904
   405
definition
huffman@25904
   406
  convex_bind_basis ::
huffman@25904
   407
  "'a pd_basis \<Rightarrow> ('a \<rightarrow> 'b convex_pd) \<rightarrow> 'b convex_pd" where
huffman@25904
   408
  "convex_bind_basis = fold_pd
huffman@25904
   409
    (\<lambda>a. \<Lambda> f. f\<cdot>(Rep_compact_basis a))
huffman@26927
   410
    (\<lambda>x y. \<Lambda> f. x\<cdot>f +\<natural> y\<cdot>f)"
huffman@25904
   411
huffman@26927
   412
lemma ACI_convex_bind:
huffman@26927
   413
  "ab_semigroup_idem_mult (\<lambda>x y. \<Lambda> f. x\<cdot>f +\<natural> y\<cdot>f)"
huffman@25904
   414
apply unfold_locales
haftmann@26041
   415
apply (simp add: convex_plus_assoc)
huffman@25904
   416
apply (simp add: convex_plus_commute)
huffman@25904
   417
apply (simp add: convex_plus_absorb eta_cfun)
huffman@25904
   418
done
huffman@25904
   419
huffman@25904
   420
lemma convex_bind_basis_simps [simp]:
huffman@25904
   421
  "convex_bind_basis (PDUnit a) =
huffman@25904
   422
    (\<Lambda> f. f\<cdot>(Rep_compact_basis a))"
huffman@25904
   423
  "convex_bind_basis (PDPlus t u) =
huffman@26927
   424
    (\<Lambda> f. convex_bind_basis t\<cdot>f +\<natural> convex_bind_basis u\<cdot>f)"
huffman@25904
   425
unfolding convex_bind_basis_def
huffman@25904
   426
apply -
huffman@26927
   427
apply (rule fold_pd_PDUnit [OF ACI_convex_bind])
huffman@26927
   428
apply (rule fold_pd_PDPlus [OF ACI_convex_bind])
huffman@25904
   429
done
huffman@25904
   430
huffman@25904
   431
lemma monofun_LAM:
huffman@25904
   432
  "\<lbrakk>cont f; cont g; \<And>x. f x \<sqsubseteq> g x\<rbrakk> \<Longrightarrow> (\<Lambda> x. f x) \<sqsubseteq> (\<Lambda> x. g x)"
huffman@25904
   433
by (simp add: expand_cfun_less)
huffman@25904
   434
huffman@25904
   435
lemma convex_bind_basis_mono:
huffman@25904
   436
  "t \<le>\<natural> u \<Longrightarrow> convex_bind_basis t \<sqsubseteq> convex_bind_basis u"
huffman@25904
   437
apply (erule convex_le_induct)
huffman@25904
   438
apply (erule (1) trans_less)
huffman@27289
   439
apply (simp add: monofun_LAM monofun_cfun)
huffman@27289
   440
apply (simp add: monofun_LAM monofun_cfun)
huffman@25904
   441
done
huffman@25904
   442
huffman@25904
   443
definition
huffman@25904
   444
  convex_bind :: "'a convex_pd \<rightarrow> ('a \<rightarrow> 'b convex_pd) \<rightarrow> 'b convex_pd" where
huffman@25904
   445
  "convex_bind = convex_pd.basis_fun convex_bind_basis"
huffman@25904
   446
huffman@25904
   447
lemma convex_bind_principal [simp]:
huffman@25904
   448
  "convex_bind\<cdot>(convex_principal t) = convex_bind_basis t"
huffman@25904
   449
unfolding convex_bind_def
huffman@25904
   450
apply (rule convex_pd.basis_fun_principal)
huffman@25904
   451
apply (erule convex_bind_basis_mono)
huffman@25904
   452
done
huffman@25904
   453
huffman@25904
   454
lemma convex_bind_unit [simp]:
huffman@26927
   455
  "convex_bind\<cdot>{x}\<natural>\<cdot>f = f\<cdot>x"
huffman@27289
   456
by (induct x rule: compact_basis.principal_induct, simp, simp)
huffman@25904
   457
huffman@25904
   458
lemma convex_bind_plus [simp]:
huffman@26927
   459
  "convex_bind\<cdot>(xs +\<natural> ys)\<cdot>f = convex_bind\<cdot>xs\<cdot>f +\<natural> convex_bind\<cdot>ys\<cdot>f"
huffman@27289
   460
by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp)
huffman@25904
   461
huffman@25904
   462
lemma convex_bind_strict [simp]: "convex_bind\<cdot>\<bottom>\<cdot>f = f\<cdot>\<bottom>"
huffman@25904
   463
unfolding convex_unit_strict [symmetric] by (rule convex_bind_unit)
huffman@25904
   464
huffman@25904
   465
huffman@25904
   466
subsection {* Map and join *}
huffman@25904
   467
huffman@25904
   468
definition
huffman@25904
   469
  convex_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a convex_pd \<rightarrow> 'b convex_pd" where
huffman@26927
   470
  "convex_map = (\<Lambda> f xs. convex_bind\<cdot>xs\<cdot>(\<Lambda> x. {f\<cdot>x}\<natural>))"
huffman@25904
   471
huffman@25904
   472
definition
huffman@25904
   473
  convex_join :: "'a convex_pd convex_pd \<rightarrow> 'a convex_pd" where
huffman@25904
   474
  "convex_join = (\<Lambda> xss. convex_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))"
huffman@25904
   475
huffman@25904
   476
lemma convex_map_unit [simp]:
huffman@25904
   477
  "convex_map\<cdot>f\<cdot>(convex_unit\<cdot>x) = convex_unit\<cdot>(f\<cdot>x)"
huffman@25904
   478
unfolding convex_map_def by simp
huffman@25904
   479
huffman@25904
   480
lemma convex_map_plus [simp]:
huffman@26927
   481
  "convex_map\<cdot>f\<cdot>(xs +\<natural> ys) = convex_map\<cdot>f\<cdot>xs +\<natural> convex_map\<cdot>f\<cdot>ys"
huffman@25904
   482
unfolding convex_map_def by simp
huffman@25904
   483
huffman@25904
   484
lemma convex_join_unit [simp]:
huffman@26927
   485
  "convex_join\<cdot>{xs}\<natural> = xs"
huffman@25904
   486
unfolding convex_join_def by simp
huffman@25904
   487
huffman@25904
   488
lemma convex_join_plus [simp]:
huffman@26927
   489
  "convex_join\<cdot>(xss +\<natural> yss) = convex_join\<cdot>xss +\<natural> convex_join\<cdot>yss"
huffman@25904
   490
unfolding convex_join_def by simp
huffman@25904
   491
huffman@25904
   492
lemma convex_map_ident: "convex_map\<cdot>(\<Lambda> x. x)\<cdot>xs = xs"
huffman@25904
   493
by (induct xs rule: convex_pd_induct, simp_all)
huffman@25904
   494
huffman@25904
   495
lemma convex_map_map:
huffman@25904
   496
  "convex_map\<cdot>f\<cdot>(convex_map\<cdot>g\<cdot>xs) = convex_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>xs"
huffman@25904
   497
by (induct xs rule: convex_pd_induct, simp_all)
huffman@25904
   498
huffman@25904
   499
lemma convex_join_map_unit:
huffman@25904
   500
  "convex_join\<cdot>(convex_map\<cdot>convex_unit\<cdot>xs) = xs"
huffman@25904
   501
by (induct xs rule: convex_pd_induct, simp_all)
huffman@25904
   502
huffman@25904
   503
lemma convex_join_map_join:
huffman@25904
   504
  "convex_join\<cdot>(convex_map\<cdot>convex_join\<cdot>xsss) = convex_join\<cdot>(convex_join\<cdot>xsss)"
huffman@25904
   505
by (induct xsss rule: convex_pd_induct, simp_all)
huffman@25904
   506
huffman@25904
   507
lemma convex_join_map_map:
huffman@25904
   508
  "convex_join\<cdot>(convex_map\<cdot>(convex_map\<cdot>f)\<cdot>xss) =
huffman@25904
   509
   convex_map\<cdot>f\<cdot>(convex_join\<cdot>xss)"
huffman@25904
   510
by (induct xss rule: convex_pd_induct, simp_all)
huffman@25904
   511
huffman@25904
   512
lemma convex_map_approx: "convex_map\<cdot>(approx n)\<cdot>xs = approx n\<cdot>xs"
huffman@25904
   513
by (induct xs rule: convex_pd_induct, simp_all)
huffman@25904
   514
huffman@25904
   515
huffman@25904
   516
subsection {* Conversions to other powerdomains *}
huffman@25904
   517
huffman@25904
   518
text {* Convex to upper *}
huffman@25904
   519
huffman@25904
   520
lemma convex_le_imp_upper_le: "t \<le>\<natural> u \<Longrightarrow> t \<le>\<sharp> u"
huffman@25904
   521
unfolding convex_le_def by simp
huffman@25904
   522
huffman@25904
   523
definition
huffman@25904
   524
  convex_to_upper :: "'a convex_pd \<rightarrow> 'a upper_pd" where
huffman@25904
   525
  "convex_to_upper = convex_pd.basis_fun upper_principal"
huffman@25904
   526
huffman@25904
   527
lemma convex_to_upper_principal [simp]:
huffman@25904
   528
  "convex_to_upper\<cdot>(convex_principal t) = upper_principal t"
huffman@25904
   529
unfolding convex_to_upper_def
huffman@25904
   530
apply (rule convex_pd.basis_fun_principal)
huffman@27289
   531
apply (rule upper_pd.principal_mono)
huffman@25904
   532
apply (erule convex_le_imp_upper_le)
huffman@25904
   533
done
huffman@25904
   534
huffman@25904
   535
lemma convex_to_upper_unit [simp]:
huffman@26927
   536
  "convex_to_upper\<cdot>{x}\<natural> = {x}\<sharp>"
huffman@27289
   537
by (induct x rule: compact_basis.principal_induct, simp, simp)
huffman@25904
   538
huffman@25904
   539
lemma convex_to_upper_plus [simp]:
huffman@26927
   540
  "convex_to_upper\<cdot>(xs +\<natural> ys) = convex_to_upper\<cdot>xs +\<sharp> convex_to_upper\<cdot>ys"
huffman@27289
   541
by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp)
huffman@25904
   542
huffman@25904
   543
lemma approx_convex_to_upper:
huffman@25904
   544
  "approx i\<cdot>(convex_to_upper\<cdot>xs) = convex_to_upper\<cdot>(approx i\<cdot>xs)"
huffman@25904
   545
by (induct xs rule: convex_pd_induct, simp, simp, simp)
huffman@25904
   546
huffman@27289
   547
lemma convex_to_upper_bind [simp]:
huffman@27289
   548
  "convex_to_upper\<cdot>(convex_bind\<cdot>xs\<cdot>f) =
huffman@27289
   549
    upper_bind\<cdot>(convex_to_upper\<cdot>xs)\<cdot>(convex_to_upper oo f)"
huffman@27289
   550
by (induct xs rule: convex_pd_induct, simp, simp, simp)
huffman@27289
   551
huffman@27289
   552
lemma convex_to_upper_map [simp]:
huffman@27289
   553
  "convex_to_upper\<cdot>(convex_map\<cdot>f\<cdot>xs) = upper_map\<cdot>f\<cdot>(convex_to_upper\<cdot>xs)"
huffman@27289
   554
by (simp add: convex_map_def upper_map_def cfcomp_LAM)
huffman@27289
   555
huffman@27289
   556
lemma convex_to_upper_join [simp]:
huffman@27289
   557
  "convex_to_upper\<cdot>(convex_join\<cdot>xss) =
huffman@27289
   558
    upper_bind\<cdot>(convex_to_upper\<cdot>xss)\<cdot>convex_to_upper"
huffman@27289
   559
by (simp add: convex_join_def upper_join_def cfcomp_LAM eta_cfun)
huffman@27289
   560
huffman@25904
   561
text {* Convex to lower *}
huffman@25904
   562
huffman@25904
   563
lemma convex_le_imp_lower_le: "t \<le>\<natural> u \<Longrightarrow> t \<le>\<flat> u"
huffman@25904
   564
unfolding convex_le_def by simp
huffman@25904
   565
huffman@25904
   566
definition
huffman@25904
   567
  convex_to_lower :: "'a convex_pd \<rightarrow> 'a lower_pd" where
huffman@25904
   568
  "convex_to_lower = convex_pd.basis_fun lower_principal"
huffman@25904
   569
huffman@25904
   570
lemma convex_to_lower_principal [simp]:
huffman@25904
   571
  "convex_to_lower\<cdot>(convex_principal t) = lower_principal t"
huffman@25904
   572
unfolding convex_to_lower_def
huffman@25904
   573
apply (rule convex_pd.basis_fun_principal)
huffman@27289
   574
apply (rule lower_pd.principal_mono)
huffman@25904
   575
apply (erule convex_le_imp_lower_le)
huffman@25904
   576
done
huffman@25904
   577
huffman@25904
   578
lemma convex_to_lower_unit [simp]:
huffman@26927
   579
  "convex_to_lower\<cdot>{x}\<natural> = {x}\<flat>"
huffman@27289
   580
by (induct x rule: compact_basis.principal_induct, simp, simp)
huffman@25904
   581
huffman@25904
   582
lemma convex_to_lower_plus [simp]:
huffman@26927
   583
  "convex_to_lower\<cdot>(xs +\<natural> ys) = convex_to_lower\<cdot>xs +\<flat> convex_to_lower\<cdot>ys"
huffman@27289
   584
by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp)
huffman@25904
   585
huffman@25904
   586
lemma approx_convex_to_lower:
huffman@25904
   587
  "approx i\<cdot>(convex_to_lower\<cdot>xs) = convex_to_lower\<cdot>(approx i\<cdot>xs)"
huffman@25904
   588
by (induct xs rule: convex_pd_induct, simp, simp, simp)
huffman@25904
   589
huffman@27289
   590
lemma convex_to_lower_bind [simp]:
huffman@27289
   591
  "convex_to_lower\<cdot>(convex_bind\<cdot>xs\<cdot>f) =
huffman@27289
   592
    lower_bind\<cdot>(convex_to_lower\<cdot>xs)\<cdot>(convex_to_lower oo f)"
huffman@27289
   593
by (induct xs rule: convex_pd_induct, simp, simp, simp)
huffman@27289
   594
huffman@27289
   595
lemma convex_to_lower_map [simp]:
huffman@27289
   596
  "convex_to_lower\<cdot>(convex_map\<cdot>f\<cdot>xs) = lower_map\<cdot>f\<cdot>(convex_to_lower\<cdot>xs)"
huffman@27289
   597
by (simp add: convex_map_def lower_map_def cfcomp_LAM)
huffman@27289
   598
huffman@27289
   599
lemma convex_to_lower_join [simp]:
huffman@27289
   600
  "convex_to_lower\<cdot>(convex_join\<cdot>xss) =
huffman@27289
   601
    lower_bind\<cdot>(convex_to_lower\<cdot>xss)\<cdot>convex_to_lower"
huffman@27289
   602
by (simp add: convex_join_def lower_join_def cfcomp_LAM eta_cfun)
huffman@27289
   603
huffman@25904
   604
text {* Ordering property *}
huffman@25904
   605
huffman@25904
   606
lemma convex_pd_less_iff:
huffman@25904
   607
  "(xs \<sqsubseteq> ys) =
huffman@25904
   608
    (convex_to_upper\<cdot>xs \<sqsubseteq> convex_to_upper\<cdot>ys \<and>
huffman@25904
   609
     convex_to_lower\<cdot>xs \<sqsubseteq> convex_to_lower\<cdot>ys)"
huffman@25904
   610
 apply (safe elim!: monofun_cfun_arg)
huffman@27309
   611
 apply (rule profinite_less_ext)
huffman@25904
   612
 apply (drule_tac f="approx i" in monofun_cfun_arg)
huffman@25904
   613
 apply (drule_tac f="approx i" in monofun_cfun_arg)
huffman@27289
   614
 apply (cut_tac x="approx i\<cdot>xs" in convex_pd.compact_imp_principal, simp)
huffman@27289
   615
 apply (cut_tac x="approx i\<cdot>ys" in convex_pd.compact_imp_principal, simp)
huffman@25904
   616
 apply clarify
huffman@25904
   617
 apply (simp add: approx_convex_to_upper approx_convex_to_lower convex_le_def)
huffman@25904
   618
done
huffman@25904
   619
huffman@26927
   620
lemmas convex_plus_less_plus_iff =
huffman@26927
   621
  convex_pd_less_iff [where xs="xs +\<natural> ys" and ys="zs +\<natural> ws", standard]
huffman@26927
   622
huffman@26927
   623
lemmas convex_pd_less_simps =
huffman@26927
   624
  convex_unit_less_plus_iff
huffman@26927
   625
  convex_plus_less_unit_iff
huffman@26927
   626
  convex_plus_less_plus_iff
huffman@26927
   627
  convex_unit_less_iff
huffman@26927
   628
  convex_to_upper_unit
huffman@26927
   629
  convex_to_upper_plus
huffman@26927
   630
  convex_to_lower_unit
huffman@26927
   631
  convex_to_lower_plus
huffman@26927
   632
  upper_pd_less_simps
huffman@26927
   633
  lower_pd_less_simps
huffman@26927
   634
huffman@25904
   635
end