diff -r 6c12f5e24e34 -r 0437dbc127b3 src/HOL/HOLCF/ConvexPD.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HOLCF/ConvexPD.thy Sat Nov 27 16:08:10 2010 -0800 @@ -0,0 +1,651 @@ +(* Title: HOLCF/ConvexPD.thy + Author: Brian Huffman +*) + +header {* Convex powerdomain *} + +theory ConvexPD +imports UpperPD LowerPD +begin + +subsection {* Basis preorder *} + +definition + convex_le :: "'a pd_basis \ 'a pd_basis \ bool" (infix "\\" 50) where + "convex_le = (\u v. u \\ v \ u \\ v)" + +lemma convex_le_refl [simp]: "t \\ t" +unfolding convex_le_def by (fast intro: upper_le_refl lower_le_refl) + +lemma convex_le_trans: "\t \\ u; u \\ v\ \ t \\ v" +unfolding convex_le_def by (fast intro: upper_le_trans lower_le_trans) + +interpretation convex_le: preorder convex_le +by (rule preorder.intro, rule convex_le_refl, rule convex_le_trans) + +lemma upper_le_minimal [simp]: "PDUnit compact_bot \\ t" +unfolding convex_le_def Rep_PDUnit by simp + +lemma PDUnit_convex_mono: "x \ y \ PDUnit x \\ PDUnit y" +unfolding convex_le_def by (fast intro: PDUnit_upper_mono PDUnit_lower_mono) + +lemma PDPlus_convex_mono: "\s \\ t; u \\ v\ \ PDPlus s u \\ PDPlus t v" +unfolding convex_le_def by (fast intro: PDPlus_upper_mono PDPlus_lower_mono) + +lemma convex_le_PDUnit_PDUnit_iff [simp]: + "(PDUnit a \\ PDUnit b) = (a \ b)" +unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit by fast + +lemma convex_le_PDUnit_lemma1: + "(PDUnit a \\ t) = (\b\Rep_pd_basis t. a \ b)" +unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit +using Rep_pd_basis_nonempty [of t, folded ex_in_conv] by fast + +lemma convex_le_PDUnit_PDPlus_iff [simp]: + "(PDUnit a \\ PDPlus t u) = (PDUnit a \\ t \ PDUnit a \\ u)" +unfolding convex_le_PDUnit_lemma1 Rep_PDPlus by fast + +lemma convex_le_PDUnit_lemma2: + "(t \\ PDUnit b) = (\a\Rep_pd_basis t. a \ b)" +unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit +using Rep_pd_basis_nonempty [of t, folded ex_in_conv] by fast + +lemma convex_le_PDPlus_PDUnit_iff [simp]: + "(PDPlus t u \\ PDUnit a) = (t \\ PDUnit a \ u \\ PDUnit a)" +unfolding convex_le_PDUnit_lemma2 Rep_PDPlus by fast + +lemma convex_le_PDPlus_lemma: + assumes z: "PDPlus t u \\ z" + shows "\v w. z = PDPlus v w \ t \\ v \ u \\ w" +proof (intro exI conjI) + let ?A = "{b\Rep_pd_basis z. \a\Rep_pd_basis t. a \ b}" + let ?B = "{b\Rep_pd_basis z. \a\Rep_pd_basis u. a \ b}" + let ?v = "Abs_pd_basis ?A" + let ?w = "Abs_pd_basis ?B" + have Rep_v: "Rep_pd_basis ?v = ?A" + apply (rule Abs_pd_basis_inverse) + apply (rule Rep_pd_basis_nonempty [of t, folded ex_in_conv, THEN exE]) + apply (cut_tac z, simp only: convex_le_def lower_le_def, clarify) + apply (drule_tac x=x in bspec, simp add: Rep_PDPlus, erule bexE) + apply (simp add: pd_basis_def) + apply fast + done + have Rep_w: "Rep_pd_basis ?w = ?B" + apply (rule Abs_pd_basis_inverse) + apply (rule Rep_pd_basis_nonempty [of u, folded ex_in_conv, THEN exE]) + apply (cut_tac z, simp only: convex_le_def lower_le_def, clarify) + apply (drule_tac x=x in bspec, simp add: Rep_PDPlus, erule bexE) + apply (simp add: pd_basis_def) + apply fast + done + show "z = PDPlus ?v ?w" + apply (insert z) + apply (simp add: convex_le_def, erule conjE) + apply (simp add: Rep_pd_basis_inject [symmetric] Rep_PDPlus) + apply (simp add: Rep_v Rep_w) + apply (rule equalityI) + apply (rule subsetI) + apply (simp only: upper_le_def) + apply (drule (1) bspec, erule bexE) + apply (simp add: Rep_PDPlus) + apply fast + apply fast + done + show "t \\ ?v" "u \\ ?w" + apply (insert z) + apply (simp_all add: convex_le_def upper_le_def lower_le_def Rep_PDPlus Rep_v Rep_w) + apply fast+ + done +qed + +lemma convex_le_induct [induct set: convex_le]: + assumes le: "t \\ u" + assumes 2: "\t u v. \P t u; P u v\ \ P t v" + assumes 3: "\a b. a \ b \ P (PDUnit a) (PDUnit b)" + assumes 4: "\t u v w. \P t v; P u w\ \ P (PDPlus t u) (PDPlus v w)" + shows "P t u" +using le apply (induct t arbitrary: u rule: pd_basis_induct) +apply (erule rev_mp) +apply (induct_tac u rule: pd_basis_induct1) +apply (simp add: 3) +apply (simp, clarify, rename_tac a b t) +apply (subgoal_tac "P (PDPlus (PDUnit a) (PDUnit a)) (PDPlus (PDUnit b) t)") +apply (simp add: PDPlus_absorb) +apply (erule (1) 4 [OF 3]) +apply (drule convex_le_PDPlus_lemma, clarify) +apply (simp add: 4) +done + + +subsection {* Type definition *} + +typedef (open) 'a convex_pd = + "{S::'a pd_basis set. convex_le.ideal S}" +by (fast intro: convex_le.ideal_principal) + +instantiation convex_pd :: ("domain") below +begin + +definition + "x \ y \ Rep_convex_pd x \ Rep_convex_pd y" + +instance .. +end + +instance convex_pd :: ("domain") po +using type_definition_convex_pd below_convex_pd_def +by (rule convex_le.typedef_ideal_po) + +instance convex_pd :: ("domain") cpo +using type_definition_convex_pd below_convex_pd_def +by (rule convex_le.typedef_ideal_cpo) + +definition + convex_principal :: "'a pd_basis \ 'a convex_pd" where + "convex_principal t = Abs_convex_pd {u. u \\ t}" + +interpretation convex_pd: + ideal_completion convex_le convex_principal Rep_convex_pd +using type_definition_convex_pd below_convex_pd_def +using convex_principal_def pd_basis_countable +by (rule convex_le.typedef_ideal_completion) + +text {* Convex powerdomain is pointed *} + +lemma convex_pd_minimal: "convex_principal (PDUnit compact_bot) \ ys" +by (induct ys rule: convex_pd.principal_induct, simp, simp) + +instance convex_pd :: ("domain") pcpo +by intro_classes (fast intro: convex_pd_minimal) + +lemma inst_convex_pd_pcpo: "\ = convex_principal (PDUnit compact_bot)" +by (rule convex_pd_minimal [THEN UU_I, symmetric]) + + +subsection {* Monadic unit and plus *} + +definition + convex_unit :: "'a \ 'a convex_pd" where + "convex_unit = compact_basis.basis_fun (\a. convex_principal (PDUnit a))" + +definition + convex_plus :: "'a convex_pd \ 'a convex_pd \ 'a convex_pd" where + "convex_plus = convex_pd.basis_fun (\t. convex_pd.basis_fun (\u. + convex_principal (PDPlus t u)))" + +abbreviation + convex_add :: "'a convex_pd \ 'a convex_pd \ 'a convex_pd" + (infixl "+\" 65) where + "xs +\ ys == convex_plus\xs\ys" + +syntax + "_convex_pd" :: "args \ 'a convex_pd" ("{_}\") + +translations + "{x,xs}\" == "{x}\ +\ {xs}\" + "{x}\" == "CONST convex_unit\x" + +lemma convex_unit_Rep_compact_basis [simp]: + "{Rep_compact_basis a}\ = convex_principal (PDUnit a)" +unfolding convex_unit_def +by (simp add: compact_basis.basis_fun_principal PDUnit_convex_mono) + +lemma convex_plus_principal [simp]: + "convex_principal t +\ convex_principal u = convex_principal (PDPlus t u)" +unfolding convex_plus_def +by (simp add: convex_pd.basis_fun_principal + convex_pd.basis_fun_mono PDPlus_convex_mono) + +interpretation convex_add: semilattice convex_add proof + fix xs ys zs :: "'a convex_pd" + show "(xs +\ ys) +\ zs = xs +\ (ys +\ zs)" + apply (induct xs ys arbitrary: zs rule: convex_pd.principal_induct2, simp, simp) + apply (rule_tac x=zs in convex_pd.principal_induct, simp) + apply (simp add: PDPlus_assoc) + done + show "xs +\ ys = ys +\ xs" + apply (induct xs ys rule: convex_pd.principal_induct2, simp, simp) + apply (simp add: PDPlus_commute) + done + show "xs +\ xs = xs" + apply (induct xs rule: convex_pd.principal_induct, simp) + apply (simp add: PDPlus_absorb) + done +qed + +lemmas convex_plus_assoc = convex_add.assoc +lemmas convex_plus_commute = convex_add.commute +lemmas convex_plus_absorb = convex_add.idem +lemmas convex_plus_left_commute = convex_add.left_commute +lemmas convex_plus_left_absorb = convex_add.left_idem + +text {* Useful for @{text "simp add: convex_plus_ac"} *} +lemmas convex_plus_ac = + convex_plus_assoc convex_plus_commute convex_plus_left_commute + +text {* Useful for @{text "simp only: convex_plus_aci"} *} +lemmas convex_plus_aci = + convex_plus_ac convex_plus_absorb convex_plus_left_absorb + +lemma convex_unit_below_plus_iff [simp]: + "{x}\ \ ys +\ zs \ {x}\ \ ys \ {x}\ \ zs" +apply (induct x rule: compact_basis.principal_induct, simp) +apply (induct ys rule: convex_pd.principal_induct, simp) +apply (induct zs rule: convex_pd.principal_induct, simp) +apply simp +done + +lemma convex_plus_below_unit_iff [simp]: + "xs +\ ys \ {z}\ \ xs \ {z}\ \ ys \ {z}\" +apply (induct xs rule: convex_pd.principal_induct, simp) +apply (induct ys rule: convex_pd.principal_induct, simp) +apply (induct z rule: compact_basis.principal_induct, simp) +apply simp +done + +lemma convex_unit_below_iff [simp]: "{x}\ \ {y}\ \ x \ y" +apply (induct x rule: compact_basis.principal_induct, simp) +apply (induct y rule: compact_basis.principal_induct, simp) +apply simp +done + +lemma convex_unit_eq_iff [simp]: "{x}\ = {y}\ \ x = y" +unfolding po_eq_conv by simp + +lemma convex_unit_strict [simp]: "{\}\ = \" +using convex_unit_Rep_compact_basis [of compact_bot] +by (simp add: inst_convex_pd_pcpo) + +lemma convex_unit_bottom_iff [simp]: "{x}\ = \ \ x = \" +unfolding convex_unit_strict [symmetric] by (rule convex_unit_eq_iff) + +lemma compact_convex_unit: "compact x \ compact {x}\" +by (auto dest!: compact_basis.compact_imp_principal) + +lemma compact_convex_unit_iff [simp]: "compact {x}\ \ compact x" +apply (safe elim!: compact_convex_unit) +apply (simp only: compact_def convex_unit_below_iff [symmetric]) +apply (erule adm_subst [OF cont_Rep_cfun2]) +done + +lemma compact_convex_plus [simp]: + "\compact xs; compact ys\ \ compact (xs +\ ys)" +by (auto dest!: convex_pd.compact_imp_principal) + + +subsection {* Induction rules *} + +lemma convex_pd_induct1: + assumes P: "adm P" + assumes unit: "\x. P {x}\" + assumes insert: "\x ys. \P {x}\; P ys\ \ P ({x}\ +\ ys)" + shows "P (xs::'a convex_pd)" +apply (induct xs rule: convex_pd.principal_induct, rule P) +apply (induct_tac a rule: pd_basis_induct1) +apply (simp only: convex_unit_Rep_compact_basis [symmetric]) +apply (rule unit) +apply (simp only: convex_unit_Rep_compact_basis [symmetric] + convex_plus_principal [symmetric]) +apply (erule insert [OF unit]) +done + +lemma convex_pd_induct + [case_names adm convex_unit convex_plus, induct type: convex_pd]: + assumes P: "adm P" + assumes unit: "\x. P {x}\" + assumes plus: "\xs ys. \P xs; P ys\ \ P (xs +\ ys)" + shows "P (xs::'a convex_pd)" +apply (induct xs rule: convex_pd.principal_induct, rule P) +apply (induct_tac a rule: pd_basis_induct) +apply (simp only: convex_unit_Rep_compact_basis [symmetric] unit) +apply (simp only: convex_plus_principal [symmetric] plus) +done + + +subsection {* Monadic bind *} + +definition + convex_bind_basis :: + "'a pd_basis \ ('a \ 'b convex_pd) \ 'b convex_pd" where + "convex_bind_basis = fold_pd + (\a. \ f. f\(Rep_compact_basis a)) + (\x y. \ f. x\f +\ y\f)" + +lemma ACI_convex_bind: + "class.ab_semigroup_idem_mult (\x y. \ f. x\f +\ y\f)" +apply unfold_locales +apply (simp add: convex_plus_assoc) +apply (simp add: convex_plus_commute) +apply (simp add: eta_cfun) +done + +lemma convex_bind_basis_simps [simp]: + "convex_bind_basis (PDUnit a) = + (\ f. f\(Rep_compact_basis a))" + "convex_bind_basis (PDPlus t u) = + (\ f. convex_bind_basis t\f +\ convex_bind_basis u\f)" +unfolding convex_bind_basis_def +apply - +apply (rule fold_pd_PDUnit [OF ACI_convex_bind]) +apply (rule fold_pd_PDPlus [OF ACI_convex_bind]) +done + +lemma convex_bind_basis_mono: + "t \\ u \ convex_bind_basis t \ convex_bind_basis u" +apply (erule convex_le_induct) +apply (erule (1) below_trans) +apply (simp add: monofun_LAM monofun_cfun) +apply (simp add: monofun_LAM monofun_cfun) +done + +definition + convex_bind :: "'a convex_pd \ ('a \ 'b convex_pd) \ 'b convex_pd" where + "convex_bind = convex_pd.basis_fun convex_bind_basis" + +lemma convex_bind_principal [simp]: + "convex_bind\(convex_principal t) = convex_bind_basis t" +unfolding convex_bind_def +apply (rule convex_pd.basis_fun_principal) +apply (erule convex_bind_basis_mono) +done + +lemma convex_bind_unit [simp]: + "convex_bind\{x}\\f = f\x" +by (induct x rule: compact_basis.principal_induct, simp, simp) + +lemma convex_bind_plus [simp]: + "convex_bind\(xs +\ ys)\f = convex_bind\xs\f +\ convex_bind\ys\f" +by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) + +lemma convex_bind_strict [simp]: "convex_bind\\\f = f\\" +unfolding convex_unit_strict [symmetric] by (rule convex_bind_unit) + +lemma convex_bind_bind: + "convex_bind\(convex_bind\xs\f)\g = + convex_bind\xs\(\ x. convex_bind\(f\x)\g)" +by (induct xs, simp_all) + + +subsection {* Map *} + +definition + convex_map :: "('a \ 'b) \ 'a convex_pd \ 'b convex_pd" where + "convex_map = (\ f xs. convex_bind\xs\(\ x. {f\x}\))" + +lemma convex_map_unit [simp]: + "convex_map\f\{x}\ = {f\x}\" +unfolding convex_map_def by simp + +lemma convex_map_plus [simp]: + "convex_map\f\(xs +\ ys) = convex_map\f\xs +\ convex_map\f\ys" +unfolding convex_map_def by simp + +lemma convex_map_bottom [simp]: "convex_map\f\\ = {f\\}\" +unfolding convex_map_def by simp + +lemma convex_map_ident: "convex_map\(\ x. x)\xs = xs" +by (induct xs rule: convex_pd_induct, simp_all) + +lemma convex_map_ID: "convex_map\ID = ID" +by (simp add: cfun_eq_iff ID_def convex_map_ident) + +lemma convex_map_map: + "convex_map\f\(convex_map\g\xs) = convex_map\(\ x. f\(g\x))\xs" +by (induct xs rule: convex_pd_induct, simp_all) + +lemma ep_pair_convex_map: "ep_pair e p \ ep_pair (convex_map\e) (convex_map\p)" +apply default +apply (induct_tac x rule: convex_pd_induct, simp_all add: ep_pair.e_inverse) +apply (induct_tac y rule: convex_pd_induct) +apply (simp_all add: ep_pair.e_p_below monofun_cfun) +done + +lemma deflation_convex_map: "deflation d \ deflation (convex_map\d)" +apply default +apply (induct_tac x rule: convex_pd_induct, simp_all add: deflation.idem) +apply (induct_tac x rule: convex_pd_induct) +apply (simp_all add: deflation.below monofun_cfun) +done + +(* FIXME: long proof! *) +lemma finite_deflation_convex_map: + assumes "finite_deflation d" shows "finite_deflation (convex_map\d)" +proof (rule finite_deflation_intro) + interpret d: finite_deflation d by fact + have "deflation d" by fact + thus "deflation (convex_map\d)" by (rule deflation_convex_map) + have "finite (range (\x. d\x))" by (rule d.finite_range) + hence "finite (Rep_compact_basis -` range (\x. d\x))" + by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject) + hence "finite (Pow (Rep_compact_basis -` range (\x. d\x)))" by simp + hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\x. d\x))))" + by (rule finite_vimageI, simp add: inj_on_def Rep_pd_basis_inject) + hence *: "finite (convex_principal ` Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\x. d\x))))" by simp + hence "finite (range (\xs. convex_map\d\xs))" + apply (rule rev_finite_subset) + apply clarsimp + apply (induct_tac xs rule: convex_pd.principal_induct) + apply (simp add: adm_mem_finite *) + apply (rename_tac t, induct_tac t rule: pd_basis_induct) + apply (simp only: convex_unit_Rep_compact_basis [symmetric] convex_map_unit) + apply simp + apply (subgoal_tac "\b. d\(Rep_compact_basis a) = Rep_compact_basis b") + apply clarsimp + apply (rule imageI) + apply (rule vimageI2) + apply (simp add: Rep_PDUnit) + apply (rule range_eqI) + apply (erule sym) + apply (rule exI) + apply (rule Abs_compact_basis_inverse [symmetric]) + apply (simp add: d.compact) + apply (simp only: convex_plus_principal [symmetric] convex_map_plus) + apply clarsimp + apply (rule imageI) + apply (rule vimageI2) + apply (simp add: Rep_PDPlus) + done + thus "finite {xs. convex_map\d\xs = xs}" + by (rule finite_range_imp_finite_fixes) +qed + +subsection {* Convex powerdomain is a domain *} + +definition + convex_approx :: "nat \ udom convex_pd \ udom convex_pd" +where + "convex_approx = (\i. convex_map\(udom_approx i))" + +lemma convex_approx: "approx_chain convex_approx" +using convex_map_ID finite_deflation_convex_map +unfolding convex_approx_def by (rule approx_chain_lemma1) + +definition convex_defl :: "defl \ defl" +where "convex_defl = defl_fun1 convex_approx convex_map" + +lemma cast_convex_defl: + "cast\(convex_defl\A) = + udom_emb convex_approx oo convex_map\(cast\A) oo udom_prj convex_approx" +using convex_approx finite_deflation_convex_map +unfolding convex_defl_def by (rule cast_defl_fun1) + +instantiation convex_pd :: ("domain") liftdomain +begin + +definition + "emb = udom_emb convex_approx oo convex_map\emb" + +definition + "prj = convex_map\prj oo udom_prj convex_approx" + +definition + "defl (t::'a convex_pd itself) = convex_defl\DEFL('a)" + +definition + "(liftemb :: 'a convex_pd u \ udom) = udom_emb u_approx oo u_map\emb" + +definition + "(liftprj :: udom \ 'a convex_pd u) = u_map\prj oo udom_prj u_approx" + +definition + "liftdefl (t::'a convex_pd itself) = u_defl\DEFL('a convex_pd)" + +instance +using liftemb_convex_pd_def liftprj_convex_pd_def liftdefl_convex_pd_def +proof (rule liftdomain_class_intro) + show "ep_pair emb (prj :: udom \ 'a convex_pd)" + unfolding emb_convex_pd_def prj_convex_pd_def + using ep_pair_udom [OF convex_approx] + by (intro ep_pair_comp ep_pair_convex_map ep_pair_emb_prj) +next + show "cast\DEFL('a convex_pd) = emb oo (prj :: udom \ 'a convex_pd)" + unfolding emb_convex_pd_def prj_convex_pd_def defl_convex_pd_def cast_convex_defl + by (simp add: cast_DEFL oo_def cfun_eq_iff convex_map_map) +qed + +end + +text {* DEFL of type constructor = type combinator *} + +lemma DEFL_convex: "DEFL('a convex_pd) = convex_defl\DEFL('a)" +by (rule defl_convex_pd_def) + + +subsection {* Join *} + +definition + convex_join :: "'a convex_pd convex_pd \ 'a convex_pd" where + "convex_join = (\ xss. convex_bind\xss\(\ xs. xs))" + +lemma convex_join_unit [simp]: + "convex_join\{xs}\ = xs" +unfolding convex_join_def by simp + +lemma convex_join_plus [simp]: + "convex_join\(xss +\ yss) = convex_join\xss +\ convex_join\yss" +unfolding convex_join_def by simp + +lemma convex_join_bottom [simp]: "convex_join\\ = \" +unfolding convex_join_def by simp + +lemma convex_join_map_unit: + "convex_join\(convex_map\convex_unit\xs) = xs" +by (induct xs rule: convex_pd_induct, simp_all) + +lemma convex_join_map_join: + "convex_join\(convex_map\convex_join\xsss) = convex_join\(convex_join\xsss)" +by (induct xsss rule: convex_pd_induct, simp_all) + +lemma convex_join_map_map: + "convex_join\(convex_map\(convex_map\f)\xss) = + convex_map\f\(convex_join\xss)" +by (induct xss rule: convex_pd_induct, simp_all) + + +subsection {* Conversions to other powerdomains *} + +text {* Convex to upper *} + +lemma convex_le_imp_upper_le: "t \\ u \ t \\ u" +unfolding convex_le_def by simp + +definition + convex_to_upper :: "'a convex_pd \ 'a upper_pd" where + "convex_to_upper = convex_pd.basis_fun upper_principal" + +lemma convex_to_upper_principal [simp]: + "convex_to_upper\(convex_principal t) = upper_principal t" +unfolding convex_to_upper_def +apply (rule convex_pd.basis_fun_principal) +apply (rule upper_pd.principal_mono) +apply (erule convex_le_imp_upper_le) +done + +lemma convex_to_upper_unit [simp]: + "convex_to_upper\{x}\ = {x}\" +by (induct x rule: compact_basis.principal_induct, simp, simp) + +lemma convex_to_upper_plus [simp]: + "convex_to_upper\(xs +\ ys) = convex_to_upper\xs +\ convex_to_upper\ys" +by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) + +lemma convex_to_upper_bind [simp]: + "convex_to_upper\(convex_bind\xs\f) = + upper_bind\(convex_to_upper\xs)\(convex_to_upper oo f)" +by (induct xs rule: convex_pd_induct, simp, simp, simp) + +lemma convex_to_upper_map [simp]: + "convex_to_upper\(convex_map\f\xs) = upper_map\f\(convex_to_upper\xs)" +by (simp add: convex_map_def upper_map_def cfcomp_LAM) + +lemma convex_to_upper_join [simp]: + "convex_to_upper\(convex_join\xss) = + upper_bind\(convex_to_upper\xss)\convex_to_upper" +by (simp add: convex_join_def upper_join_def cfcomp_LAM eta_cfun) + +text {* Convex to lower *} + +lemma convex_le_imp_lower_le: "t \\ u \ t \\ u" +unfolding convex_le_def by simp + +definition + convex_to_lower :: "'a convex_pd \ 'a lower_pd" where + "convex_to_lower = convex_pd.basis_fun lower_principal" + +lemma convex_to_lower_principal [simp]: + "convex_to_lower\(convex_principal t) = lower_principal t" +unfolding convex_to_lower_def +apply (rule convex_pd.basis_fun_principal) +apply (rule lower_pd.principal_mono) +apply (erule convex_le_imp_lower_le) +done + +lemma convex_to_lower_unit [simp]: + "convex_to_lower\{x}\ = {x}\" +by (induct x rule: compact_basis.principal_induct, simp, simp) + +lemma convex_to_lower_plus [simp]: + "convex_to_lower\(xs +\ ys) = convex_to_lower\xs +\ convex_to_lower\ys" +by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) + +lemma convex_to_lower_bind [simp]: + "convex_to_lower\(convex_bind\xs\f) = + lower_bind\(convex_to_lower\xs)\(convex_to_lower oo f)" +by (induct xs rule: convex_pd_induct, simp, simp, simp) + +lemma convex_to_lower_map [simp]: + "convex_to_lower\(convex_map\f\xs) = lower_map\f\(convex_to_lower\xs)" +by (simp add: convex_map_def lower_map_def cfcomp_LAM) + +lemma convex_to_lower_join [simp]: + "convex_to_lower\(convex_join\xss) = + lower_bind\(convex_to_lower\xss)\convex_to_lower" +by (simp add: convex_join_def lower_join_def cfcomp_LAM eta_cfun) + +text {* Ordering property *} + +lemma convex_pd_below_iff: + "(xs \ ys) = + (convex_to_upper\xs \ convex_to_upper\ys \ + convex_to_lower\xs \ convex_to_lower\ys)" +apply (induct xs rule: convex_pd.principal_induct, simp) +apply (induct ys rule: convex_pd.principal_induct, simp) +apply (simp add: convex_le_def) +done + +lemmas convex_plus_below_plus_iff = + convex_pd_below_iff [where xs="xs +\ ys" and ys="zs +\ ws", standard] + +lemmas convex_pd_below_simps = + convex_unit_below_plus_iff + convex_plus_below_unit_iff + convex_plus_below_plus_iff + convex_unit_below_iff + convex_to_upper_unit + convex_to_upper_plus + convex_to_lower_unit + convex_to_lower_plus + upper_pd_below_simps + lower_pd_below_simps + +end