Theory Directed_Completeness
theory Directed_Completeness
imports
Complete_Non_Orders.Continuity
Well_Order_Connection
"HOL-Cardinals.Cardinals"
"HOL-Library.FuncSet"
begin
subsection ‹Missing Lemmas›
no_notation disj (infixr ‹|› 30)
lemma Sup_funpow_mono:
fixes f :: "'a :: complete_lattice ⇒ 'a"
assumes mono: "mono f"
shows "mono (⨆i. f ^^ i)"
by (intro monoI, auto intro!: Sup_mono dest: funpow_mono[OF mono])
lemma iso_imp_compat:
assumes iso: "iso r r' f" shows "compat r r' f"
by (simp add: compat_def iso iso_forward)
lemma iso_inv_into:
assumes ISO: "iso r r' f"
shows "iso r' r (inv_into (Field r) f)"
using assms unfolding iso_def
using bij_betw_inv_into inv_into_Field_embed_bij_betw by blast
lemmas iso_imp_compat_inv_into = iso_imp_compat[OF iso_inv_into]
lemma infinite_iff_natLeq: "infinite A ⟷ natLeq ≤o |A|"
using infinite_iff_natLeq_ordLeq by blast
text ‹As we cannot formalize transfinite sequences directly, we take the following approach:
We just use @{term A} as the index set, and instead of the ordering on ordinals,
we take the well-order that is chosen by the cardinality library to denote @{term "|A|"}.›
definition well_order_of (‹('(≼⇘_⇙'))› [0]1000) where "(≼⇘A⇙) x y ≡ (x,y) ∈ |A|"
abbreviation well_order_le (‹_ ≼⇘_⇙ _› [51,0,51]50) where "x ≼⇘A⇙ y ≡ (≼⇘A⇙) x y"
abbreviation well_order_less (‹_ ≺⇘_⇙ _› [51,0,51]50) where "x ≺⇘A⇙ y ≡ asympartp (≼⇘A⇙) x y"
lemmas well_order_ofI = well_order_of_def[unfolded atomize_eq, THEN iffD2]
lemmas well_order_ofD = well_order_of_def[unfolded atomize_eq, THEN iffD1]
lemma carrier: assumes "x ≼⇘A⇙ y" shows "x ∈ A" and "y ∈ A"
using assms by (auto dest!: well_order_ofD dest: FieldI1 FieldI2)
lemma relation_of[simp]: "relation_of (≼⇘A⇙) A = |A|"
by (auto simp: relation_of_def well_order_of_def dest: FieldI1 FieldI2)