(* Title: HOLCF/LowerPD.thy ID: $Id: LowerPD.thy,v 1.8 2008/05/19 21:49:21 huffman Exp $ Author: Brian Huffman *) header {* Lower powerdomain *} theory LowerPD imports CompactBasis begin subsection {* Basis preorder *} definition lower_le :: "'a pd_basis => 'a pd_basis => bool" (infix "≤\<flat>" 50) where "lower_le = (λu v. ∀x∈Rep_pd_basis u. ∃y∈Rep_pd_basis v. x \<sqsubseteq> y)" lemma lower_le_refl [simp]: "t ≤\<flat> t" unfolding lower_le_def by fast lemma lower_le_trans: "[|t ≤\<flat> u; u ≤\<flat> v|] ==> t ≤\<flat> v" unfolding lower_le_def apply (rule ballI) apply (drule (1) bspec, erule bexE) apply (drule (1) bspec, erule bexE) apply (erule rev_bexI) apply (erule (1) trans_less) done interpretation lower_le: preorder [lower_le] by (rule preorder.intro, rule lower_le_refl, rule lower_le_trans) lemma lower_le_minimal [simp]: "PDUnit compact_bot ≤\<flat> t" unfolding lower_le_def Rep_PDUnit by (simp, rule Rep_pd_basis_nonempty [folded ex_in_conv]) lemma PDUnit_lower_mono: "x \<sqsubseteq> y ==> PDUnit x ≤\<flat> PDUnit y" unfolding lower_le_def Rep_PDUnit by fast lemma PDPlus_lower_mono: "[|s ≤\<flat> t; u ≤\<flat> v|] ==> PDPlus s u ≤\<flat> PDPlus t v" unfolding lower_le_def Rep_PDPlus by fast lemma PDPlus_lower_less: "t ≤\<flat> PDPlus t u" unfolding lower_le_def Rep_PDPlus by fast lemma lower_le_PDUnit_PDUnit_iff [simp]: "(PDUnit a ≤\<flat> PDUnit b) = a \<sqsubseteq> b" unfolding lower_le_def Rep_PDUnit by fast lemma lower_le_PDUnit_PDPlus_iff: "(PDUnit a ≤\<flat> PDPlus t u) = (PDUnit a ≤\<flat> t ∨ PDUnit a ≤\<flat> u)" unfolding lower_le_def Rep_PDPlus Rep_PDUnit by fast lemma lower_le_PDPlus_iff: "(PDPlus t u ≤\<flat> v) = (t ≤\<flat> v ∧ u ≤\<flat> v)" unfolding lower_le_def Rep_PDPlus by fast lemma lower_le_induct [induct set: lower_le]: assumes le: "t ≤\<flat> u" assumes 1: "!!a b. a \<sqsubseteq> b ==> P (PDUnit a) (PDUnit b)" assumes 2: "!!t u a. P (PDUnit a) t ==> P (PDUnit a) (PDPlus t u)" assumes 3: "!!t u v. [|P t v; P u v|] ==> P (PDPlus t u) v" 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_induct) apply (simp add: 1) apply (simp add: lower_le_PDUnit_PDPlus_iff) apply (simp add: 2) apply (subst PDPlus_commute) apply (simp add: 2) apply (simp add: lower_le_PDPlus_iff 3) done lemma approx_pd_lower_mono1: "i ≤ j ==> approx_pd i t ≤\<flat> approx_pd j t" apply (induct t rule: pd_basis_induct) apply (simp add: compact_approx_mono1) apply (simp add: PDPlus_lower_mono) done lemma approx_pd_lower_le: "approx_pd i t ≤\<flat> t" apply (induct t rule: pd_basis_induct) apply (simp add: compact_approx_le) apply (simp add: PDPlus_lower_mono) done lemma approx_pd_lower_mono: "t ≤\<flat> u ==> approx_pd n t ≤\<flat> approx_pd n u" apply (erule lower_le_induct) apply (simp add: compact_approx_mono) apply (simp add: lower_le_PDUnit_PDPlus_iff) apply (simp add: lower_le_PDPlus_iff) done subsection {* Type definition *} cpodef (open) 'a lower_pd = "{S::'a::profinite pd_basis set. lower_le.ideal S}" apply (simp add: lower_le.adm_ideal) apply (fast intro: lower_le.ideal_principal) done lemma ideal_Rep_lower_pd: "lower_le.ideal (Rep_lower_pd x)" by (rule Rep_lower_pd [unfolded mem_Collect_eq]) definition lower_principal :: "'a pd_basis => 'a lower_pd" where "lower_principal t = Abs_lower_pd {u. u ≤\<flat> t}" lemma Rep_lower_principal: "Rep_lower_pd (lower_principal t) = {u. u ≤\<flat> t}" unfolding lower_principal_def apply (rule Abs_lower_pd_inverse [simplified]) apply (rule lower_le.ideal_principal) done interpretation lower_pd: ideal_completion [lower_le approx_pd lower_principal Rep_lower_pd] apply unfold_locales apply (rule approx_pd_lower_le) apply (rule approx_pd_idem) apply (erule approx_pd_lower_mono) apply (rule approx_pd_lower_mono1, simp) apply (rule finite_range_approx_pd) apply (rule ex_approx_pd_eq) apply (rule ideal_Rep_lower_pd) apply (rule cont_Rep_lower_pd) apply (rule Rep_lower_principal) apply (simp only: less_lower_pd_def less_set_eq) done lemma lower_principal_less_iff [simp]: "lower_principal t \<sqsubseteq> lower_principal u <-> t ≤\<flat> u" by (rule lower_pd.principal_less_iff) lemma lower_principal_eq_iff: "lower_principal t = lower_principal u <-> t ≤\<flat> u ∧ u ≤\<flat> t" by (rule lower_pd.principal_eq_iff) lemma lower_principal_mono: "t ≤\<flat> u ==> lower_principal t \<sqsubseteq> lower_principal u" by (rule lower_pd.principal_mono) lemma compact_lower_principal: "compact (lower_principal t)" by (rule lower_pd.compact_principal) lemma lower_pd_minimal: "lower_principal (PDUnit compact_bot) \<sqsubseteq> ys" by (induct ys rule: lower_pd.principal_induct, simp, simp) instance lower_pd :: (bifinite) pcpo by intro_classes (fast intro: lower_pd_minimal) lemma inst_lower_pd_pcpo: "⊥ = lower_principal (PDUnit compact_bot)" by (rule lower_pd_minimal [THEN UU_I, symmetric]) subsection {* Approximation *} instantiation lower_pd :: (profinite) profinite begin definition approx_lower_pd_def: "approx = lower_pd.completion_approx" instance apply (intro_classes, unfold approx_lower_pd_def) apply (simp add: lower_pd.chain_completion_approx) apply (rule lower_pd.lub_completion_approx) apply (rule lower_pd.completion_approx_idem) apply (rule lower_pd.finite_fixes_completion_approx) done end instance lower_pd :: (bifinite) bifinite .. lemma approx_lower_principal [simp]: "approx n·(lower_principal t) = lower_principal (approx_pd n t)" unfolding approx_lower_pd_def by (rule lower_pd.completion_approx_principal) lemma approx_eq_lower_principal: "∃t∈Rep_lower_pd xs. approx n·xs = lower_principal (approx_pd n t)" unfolding approx_lower_pd_def by (rule lower_pd.completion_approx_eq_principal) lemma compact_imp_lower_principal: "compact xs ==> ∃t. xs = lower_principal t" apply (drule bifinite_compact_eq_approx) apply (erule exE) apply (erule subst) apply (cut_tac n=i and xs=xs in approx_eq_lower_principal) apply fast done lemma lower_principal_induct: "[|adm P; !!t. P (lower_principal t)|] ==> P xs" by (rule lower_pd.principal_induct) lemma lower_principal_induct2: "[|!!ys. adm (λxs. P xs ys); !!xs. adm (λys. P xs ys); !!t u. P (lower_principal t) (lower_principal u)|] ==> P xs ys" apply (rule_tac x=ys in spec) apply (rule_tac xs=xs in lower_principal_induct, simp) apply (rule allI, rename_tac ys) apply (rule_tac xs=ys in lower_principal_induct, simp) apply simp done subsection {* Monadic unit and plus *} definition lower_unit :: "'a -> 'a lower_pd" where "lower_unit = compact_basis.basis_fun (λa. lower_principal (PDUnit a))" definition lower_plus :: "'a lower_pd -> 'a lower_pd -> 'a lower_pd" where "lower_plus = lower_pd.basis_fun (λt. lower_pd.basis_fun (λu. lower_principal (PDPlus t u)))" abbreviation lower_add :: "'a lower_pd => 'a lower_pd => 'a lower_pd" (infixl "+\<flat>" 65) where "xs +\<flat> ys == lower_plus·xs·ys" syntax "_lower_pd" :: "args => 'a lower_pd" ("{_}\<flat>") translations "{x,xs}\<flat>" == "{x}\<flat> +\<flat> {xs}\<flat>" "{x}\<flat>" == "CONST lower_unit·x" lemma lower_unit_Rep_compact_basis [simp]: "{Rep_compact_basis a}\<flat> = lower_principal (PDUnit a)" unfolding lower_unit_def by (simp add: compact_basis.basis_fun_principal lower_principal_mono PDUnit_lower_mono) lemma lower_plus_principal [simp]: "lower_principal t +\<flat> lower_principal u = lower_principal (PDPlus t u)" unfolding lower_plus_def by (simp add: lower_pd.basis_fun_principal lower_pd.basis_fun_mono PDPlus_lower_mono) lemma approx_lower_unit [simp]: "approx n·{x}\<flat> = {approx n·x}\<flat>" apply (induct x rule: compact_basis_induct, simp) apply (simp add: approx_Rep_compact_basis) done lemma approx_lower_plus [simp]: "approx n·(xs +\<flat> ys) = (approx n·xs) +\<flat> (approx n·ys)" by (induct xs ys rule: lower_principal_induct2, simp, simp, simp) lemma lower_plus_assoc: "(xs +\<flat> ys) +\<flat> zs = xs +\<flat> (ys +\<flat> zs)" apply (induct xs ys arbitrary: zs rule: lower_principal_induct2, simp, simp) apply (rule_tac xs=zs in lower_principal_induct, simp) apply (simp add: PDPlus_assoc) done lemma lower_plus_commute: "xs +\<flat> ys = ys +\<flat> xs" apply (induct xs ys rule: lower_principal_induct2, simp, simp) apply (simp add: PDPlus_commute) done lemma lower_plus_absorb: "xs +\<flat> xs = xs" apply (induct xs rule: lower_principal_induct, simp) apply (simp add: PDPlus_absorb) done interpretation aci_lower_plus: ab_semigroup_idem_mult ["op +\<flat>"] by unfold_locales (rule lower_plus_assoc lower_plus_commute lower_plus_absorb)+ lemma lower_plus_left_commute: "xs +\<flat> (ys +\<flat> zs) = ys +\<flat> (xs +\<flat> zs)" by (rule aci_lower_plus.mult_left_commute) lemma lower_plus_left_absorb: "xs +\<flat> (xs +\<flat> ys) = xs +\<flat> ys" by (rule aci_lower_plus.mult_left_idem) lemmas lower_plus_aci = aci_lower_plus.mult_ac_idem lemma lower_plus_less1: "xs \<sqsubseteq> xs +\<flat> ys" apply (induct xs ys rule: lower_principal_induct2, simp, simp) apply (simp add: PDPlus_lower_less) done lemma lower_plus_less2: "ys \<sqsubseteq> xs +\<flat> ys" by (subst lower_plus_commute, rule lower_plus_less1) lemma lower_plus_least: "[|xs \<sqsubseteq> zs; ys \<sqsubseteq> zs|] ==> xs +\<flat> ys \<sqsubseteq> zs" apply (subst lower_plus_absorb [of zs, symmetric]) apply (erule (1) monofun_cfun [OF monofun_cfun_arg]) done lemma lower_plus_less_iff: "xs +\<flat> ys \<sqsubseteq> zs <-> xs \<sqsubseteq> zs ∧ ys \<sqsubseteq> zs" apply safe apply (erule trans_less [OF lower_plus_less1]) apply (erule trans_less [OF lower_plus_less2]) apply (erule (1) lower_plus_least) done lemma lower_unit_less_plus_iff: "{x}\<flat> \<sqsubseteq> ys +\<flat> zs <-> {x}\<flat> \<sqsubseteq> ys ∨ {x}\<flat> \<sqsubseteq> zs" apply (rule iffI) apply (subgoal_tac "adm (λf. f·{x}\<flat> \<sqsubseteq> f·ys ∨ f·{x}\<flat> \<sqsubseteq> f·zs)") apply (drule admD, rule chain_approx) apply (drule_tac f="approx i" in monofun_cfun_arg) apply (cut_tac x="approx i·x" in compact_imp_Rep_compact_basis, simp) apply (cut_tac xs="approx i·ys" in compact_imp_lower_principal, simp) apply (cut_tac xs="approx i·zs" in compact_imp_lower_principal, simp) apply (clarify, simp add: lower_le_PDUnit_PDPlus_iff) apply simp apply simp apply (erule disjE) apply (erule trans_less [OF _ lower_plus_less1]) apply (erule trans_less [OF _ lower_plus_less2]) done lemma lower_unit_less_iff [simp]: "{x}\<flat> \<sqsubseteq> {y}\<flat> <-> x \<sqsubseteq> y" apply (rule iffI) apply (rule bifinite_less_ext) apply (drule_tac f="approx i" in monofun_cfun_arg, simp) apply (cut_tac x="approx i·x" in compact_imp_Rep_compact_basis, simp) apply (cut_tac x="approx i·y" in compact_imp_Rep_compact_basis, simp) apply (clarify, simp add: compact_le_def) apply (erule monofun_cfun_arg) done lemmas lower_pd_less_simps = lower_unit_less_iff lower_plus_less_iff lower_unit_less_plus_iff lemma lower_unit_eq_iff [simp]: "{x}\<flat> = {y}\<flat> <-> x = y" unfolding po_eq_conv by simp lemma lower_unit_strict [simp]: "{⊥}\<flat> = ⊥" unfolding inst_lower_pd_pcpo Rep_compact_bot [symmetric] by simp lemma lower_unit_strict_iff [simp]: "{x}\<flat> = ⊥ <-> x = ⊥" unfolding lower_unit_strict [symmetric] by (rule lower_unit_eq_iff) lemma lower_plus_strict_iff [simp]: "xs +\<flat> ys = ⊥ <-> xs = ⊥ ∧ ys = ⊥" apply safe apply (rule UU_I, erule subst, rule lower_plus_less1) apply (rule UU_I, erule subst, rule lower_plus_less2) apply (rule lower_plus_absorb) done lemma lower_plus_strict1 [simp]: "⊥ +\<flat> ys = ys" apply (rule antisym_less [OF _ lower_plus_less2]) apply (simp add: lower_plus_least) done lemma lower_plus_strict2 [simp]: "xs +\<flat> ⊥ = xs" apply (rule antisym_less [OF _ lower_plus_less1]) apply (simp add: lower_plus_least) done lemma compact_lower_unit_iff [simp]: "compact {x}\<flat> <-> compact x" unfolding bifinite_compact_iff by simp lemma compact_lower_plus [simp]: "[|compact xs; compact ys|] ==> compact (xs +\<flat> ys)" apply (drule compact_imp_lower_principal)+ apply (auto simp add: compact_lower_principal) done subsection {* Induction rules *} lemma lower_pd_induct1: assumes P: "adm P" assumes unit: "!!x. P {x}\<flat>" assumes insert: "!!x ys. [|P {x}\<flat>; P ys|] ==> P ({x}\<flat> +\<flat> ys)" shows "P (xs::'a lower_pd)" apply (induct xs rule: lower_principal_induct, rule P) apply (induct_tac t rule: pd_basis_induct1) apply (simp only: lower_unit_Rep_compact_basis [symmetric]) apply (rule unit) apply (simp only: lower_unit_Rep_compact_basis [symmetric] lower_plus_principal [symmetric]) apply (erule insert [OF unit]) done lemma lower_pd_induct: assumes P: "adm P" assumes unit: "!!x. P {x}\<flat>" assumes plus: "!!xs ys. [|P xs; P ys|] ==> P (xs +\<flat> ys)" shows "P (xs::'a lower_pd)" apply (induct xs rule: lower_principal_induct, rule P) apply (induct_tac t rule: pd_basis_induct) apply (simp only: lower_unit_Rep_compact_basis [symmetric] unit) apply (simp only: lower_plus_principal [symmetric] plus) done subsection {* Monadic bind *} definition lower_bind_basis :: "'a pd_basis => ('a -> 'b lower_pd) -> 'b lower_pd" where "lower_bind_basis = fold_pd (λa. Λ f. f·(Rep_compact_basis a)) (λx y. Λ f. x·f +\<flat> y·f)" lemma ACI_lower_bind: "ab_semigroup_idem_mult (λx y. Λ f. x·f +\<flat> y·f)" apply unfold_locales apply (simp add: lower_plus_assoc) apply (simp add: lower_plus_commute) apply (simp add: lower_plus_absorb eta_cfun) done lemma lower_bind_basis_simps [simp]: "lower_bind_basis (PDUnit a) = (Λ f. f·(Rep_compact_basis a))" "lower_bind_basis (PDPlus t u) = (Λ f. lower_bind_basis t·f +\<flat> lower_bind_basis u·f)" unfolding lower_bind_basis_def apply - apply (rule fold_pd_PDUnit [OF ACI_lower_bind]) apply (rule fold_pd_PDPlus [OF ACI_lower_bind]) done lemma lower_bind_basis_mono: "t ≤\<flat> u ==> lower_bind_basis t \<sqsubseteq> lower_bind_basis u" unfolding expand_cfun_less apply (erule lower_le_induct, safe) apply (simp add: compact_le_def monofun_cfun) apply (simp add: rev_trans_less [OF lower_plus_less1]) apply (simp add: lower_plus_less_iff) done definition lower_bind :: "'a lower_pd -> ('a -> 'b lower_pd) -> 'b lower_pd" where "lower_bind = lower_pd.basis_fun lower_bind_basis" lemma lower_bind_principal [simp]: "lower_bind·(lower_principal t) = lower_bind_basis t" unfolding lower_bind_def apply (rule lower_pd.basis_fun_principal) apply (erule lower_bind_basis_mono) done lemma lower_bind_unit [simp]: "lower_bind·{x}\<flat>·f = f·x" by (induct x rule: compact_basis_induct, simp, simp) lemma lower_bind_plus [simp]: "lower_bind·(xs +\<flat> ys)·f = lower_bind·xs·f +\<flat> lower_bind·ys·f" by (induct xs ys rule: lower_principal_induct2, simp, simp, simp) lemma lower_bind_strict [simp]: "lower_bind·⊥·f = f·⊥" unfolding lower_unit_strict [symmetric] by (rule lower_bind_unit) subsection {* Map and join *} definition lower_map :: "('a -> 'b) -> 'a lower_pd -> 'b lower_pd" where "lower_map = (Λ f xs. lower_bind·xs·(Λ x. {f·x}\<flat>))" definition lower_join :: "'a lower_pd lower_pd -> 'a lower_pd" where "lower_join = (Λ xss. lower_bind·xss·(Λ xs. xs))" lemma lower_map_unit [simp]: "lower_map·f·{x}\<flat> = {f·x}\<flat>" unfolding lower_map_def by simp lemma lower_map_plus [simp]: "lower_map·f·(xs +\<flat> ys) = lower_map·f·xs +\<flat> lower_map·f·ys" unfolding lower_map_def by simp lemma lower_join_unit [simp]: "lower_join·{xs}\<flat> = xs" unfolding lower_join_def by simp lemma lower_join_plus [simp]: "lower_join·(xss +\<flat> yss) = lower_join·xss +\<flat> lower_join·yss" unfolding lower_join_def by simp lemma lower_map_ident: "lower_map·(Λ x. x)·xs = xs" by (induct xs rule: lower_pd_induct, simp_all) lemma lower_map_map: "lower_map·f·(lower_map·g·xs) = lower_map·(Λ x. f·(g·x))·xs" by (induct xs rule: lower_pd_induct, simp_all) lemma lower_join_map_unit: "lower_join·(lower_map·lower_unit·xs) = xs" by (induct xs rule: lower_pd_induct, simp_all) lemma lower_join_map_join: "lower_join·(lower_map·lower_join·xsss) = lower_join·(lower_join·xsss)" by (induct xsss rule: lower_pd_induct, simp_all) lemma lower_join_map_map: "lower_join·(lower_map·(lower_map·f)·xss) = lower_map·f·(lower_join·xss)" by (induct xss rule: lower_pd_induct, simp_all) lemma lower_map_approx: "lower_map·(approx n)·xs = approx n·xs" by (induct xs rule: lower_pd_induct, simp_all) end
lemma lower_le_refl:
t ≤\<flat> t
lemma lower_le_trans:
[| t ≤\<flat> u; u ≤\<flat> v |] ==> t ≤\<flat> v
lemma lower_le_minimal:
PDUnit compact_bot ≤\<flat> t
lemma PDUnit_lower_mono:
x << y ==> PDUnit x ≤\<flat> PDUnit y
lemma PDPlus_lower_mono:
[| s ≤\<flat> t; u ≤\<flat> v |] ==> PDPlus s u ≤\<flat> PDPlus t v
lemma PDPlus_lower_less:
t ≤\<flat> PDPlus t u
lemma lower_le_PDUnit_PDUnit_iff:
(PDUnit a ≤\<flat> PDUnit b) = a << b
lemma lower_le_PDUnit_PDPlus_iff:
(PDUnit a ≤\<flat> PDPlus t u) = (PDUnit a ≤\<flat> t ∨ PDUnit a ≤\<flat> u)
lemma lower_le_PDPlus_iff:
(PDPlus t u ≤\<flat> v) = (t ≤\<flat> v ∧ u ≤\<flat> v)
lemma lower_le_induct:
[| t ≤\<flat> u; !!a b. a << b ==> P (PDUnit a) (PDUnit b);
!!t u a. P (PDUnit a) t ==> P (PDUnit a) (PDPlus t u);
!!t u v. [| P t v; P u v |] ==> P (PDPlus t u) v |]
==> P t u
lemma approx_pd_lower_mono1:
i ≤ j ==> approx_pd i t ≤\<flat> approx_pd j t
lemma approx_pd_lower_le:
approx_pd i t ≤\<flat> t
lemma approx_pd_lower_mono:
t ≤\<flat> u ==> approx_pd n t ≤\<flat> approx_pd n u
lemma ideal_Rep_lower_pd:
preorder.ideal op ≤\<flat> (Rep_lower_pd x)
lemma Rep_lower_principal:
Rep_lower_pd (lower_principal t) = {u. u ≤\<flat> t}
lemma lower_principal_less_iff:
lower_principal t << lower_principal u = (t ≤\<flat> u)
lemma lower_principal_eq_iff:
(lower_principal t = lower_principal u) = (t ≤\<flat> u ∧ u ≤\<flat> t)
lemma lower_principal_mono:
t ≤\<flat> u ==> lower_principal t << lower_principal u
lemma compact_lower_principal:
compact (lower_principal t)
lemma lower_pd_minimal:
lower_principal (PDUnit compact_bot) << ys
lemma inst_lower_pd_pcpo:
UU = lower_principal (PDUnit compact_bot)
lemma approx_lower_principal:
approx n·(lower_principal t) = lower_principal (approx_pd n t)
lemma approx_eq_lower_principal:
∃t∈Rep_lower_pd xs. approx n·xs = lower_principal (approx_pd n t)
lemma compact_imp_lower_principal:
compact xs ==> ∃t. xs = lower_principal t
lemma lower_principal_induct:
[| adm P; !!t. P (lower_principal t) |] ==> P xs
lemma lower_principal_induct2:
[| !!ys. adm (λxs. P xs ys); !!xs. adm (P xs);
!!t u. P (lower_principal t) (lower_principal u) |]
==> P xs ys
lemma lower_unit_Rep_compact_basis:
{Rep_compact_basis a}\<flat> = lower_principal (PDUnit a)
lemma lower_plus_principal:
lower_principal t +\<flat> lower_principal u = lower_principal (PDPlus t u)
lemma approx_lower_unit:
approx n·{x}\<flat> = {approx n·x}\<flat>
lemma approx_lower_plus:
approx n·(xs +\<flat> ys) = approx n·xs +\<flat> approx n·ys
lemma lower_plus_assoc:
xs +\<flat> ys +\<flat> zs = xs +\<flat> (ys +\<flat> zs)
lemma lower_plus_commute:
xs +\<flat> ys = ys +\<flat> xs
lemma lower_plus_absorb:
xs +\<flat> xs = xs
lemma lower_plus_left_commute:
xs +\<flat> (ys +\<flat> zs) = ys +\<flat> (xs +\<flat> zs)
lemma lower_plus_left_absorb:
xs +\<flat> (xs +\<flat> ys) = xs +\<flat> ys
lemma lower_plus_aci:
a +\<flat> b +\<flat> c = a +\<flat> (b +\<flat> c)
a +\<flat> b = b +\<flat> a
a +\<flat> (b +\<flat> c) = b +\<flat> (a +\<flat> c)
x +\<flat> x = x
x +\<flat> (x +\<flat> y) = x +\<flat> y
lemma lower_plus_less1:
xs << xs +\<flat> ys
lemma lower_plus_less2:
ys << xs +\<flat> ys
lemma lower_plus_least:
[| xs << zs; ys << zs |] ==> xs +\<flat> ys << zs
lemma lower_plus_less_iff:
xs +\<flat> ys << zs = (xs << zs ∧ ys << zs)
lemma lower_unit_less_plus_iff:
{x}\<flat> << ys +\<flat> zs = ({x}\<flat> << ys ∨ {x}\<flat> << zs)
lemma lower_unit_less_iff:
{x}\<flat> << {y}\<flat> = x << y
lemma lower_pd_less_simps:
{x}\<flat> << {y}\<flat> = x << y
xs +\<flat> ys << zs = (xs << zs ∧ ys << zs)
{x}\<flat> << ys +\<flat> zs = ({x}\<flat> << ys ∨ {x}\<flat> << zs)
lemma lower_unit_eq_iff:
({x}\<flat> = {y}\<flat>) = (x = y)
lemma lower_unit_strict:
{UU}\<flat> = UU
lemma lower_unit_strict_iff:
({x}\<flat> = UU) = (x = UU)
lemma lower_plus_strict_iff:
(xs +\<flat> ys = UU) = (xs = UU ∧ ys = UU)
lemma lower_plus_strict1:
UU +\<flat> ys = ys
lemma lower_plus_strict2:
xs +\<flat> UU = xs
lemma compact_lower_unit_iff:
compact {x}\<flat> = compact x
lemma compact_lower_plus:
[| compact xs; compact ys |] ==> compact (xs +\<flat> ys)
lemma lower_pd_induct1:
[| adm P; !!x. P {x}\<flat>;
!!x ys. [| P {x}\<flat>; P ys |] ==> P ({x}\<flat> +\<flat> ys) |]
==> P xs
lemma lower_pd_induct:
[| adm P; !!x. P {x}\<flat>; !!xs ys. [| P xs; P ys |] ==> P (xs +\<flat> ys) |]
==> P xs
lemma ACI_lower_bind:
ab_semigroup_idem_mult (λx y. LAM f. x·f +\<flat> y·f)
lemma lower_bind_basis_simps:
lower_bind_basis (PDUnit a) = (LAM f. f·(Rep_compact_basis a))
lower_bind_basis (PDPlus t u) =
(LAM f. lower_bind_basis t·f +\<flat> lower_bind_basis u·f)
lemma lower_bind_basis_mono:
t ≤\<flat> u ==> lower_bind_basis t << lower_bind_basis u
lemma lower_bind_principal:
lower_bind·(lower_principal t) = lower_bind_basis t
lemma lower_bind_unit:
lower_bind·{x}\<flat>·f = f·x
lemma lower_bind_plus:
lower_bind·(xs +\<flat> ys)·f = lower_bind·xs·f +\<flat> lower_bind·ys·f
lemma lower_bind_strict:
lower_bind·UU·f = f·UU
lemma lower_map_unit:
lower_map·f·{x}\<flat> = {f·x}\<flat>
lemma lower_map_plus:
lower_map·f·(xs +\<flat> ys) = lower_map·f·xs +\<flat> lower_map·f·ys
lemma lower_join_unit:
lower_join·{xs}\<flat> = xs
lemma lower_join_plus:
lower_join·(xss +\<flat> yss) = lower_join·xss +\<flat> lower_join·yss
lemma lower_map_ident:
lower_map·(LAM x. x)·xs = xs
lemma lower_map_map:
lower_map·f·(lower_map·g·xs) = lower_map·(LAM x. f·(g·x))·xs
lemma lower_join_map_unit:
lower_join·(lower_map·lower_unit·xs) = xs
lemma lower_join_map_join:
lower_join·(lower_map·lower_join·xsss) = lower_join·(lower_join·xsss)
lemma lower_join_map_map:
lower_join·(lower_map·(lower_map·f)·xss) = lower_map·f·(lower_join·xss)
lemma lower_map_approx:
lower_map·(approx n)·xs = approx n·xs