--- a/src/HOL/HOLCF/Adm.thy Tue Dec 10 22:40:07 2024 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,179 +0,0 @@
-(* Title: HOL/HOLCF/Adm.thy
- Author: Franz Regensburger and Brian Huffman
-*)
-
-section \<open>Admissibility and compactness\<close>
-
-theory Adm
- imports Cont
-begin
-
-default_sort cpo
-
-subsection \<open>Definitions\<close>
-
-definition adm :: "('a::cpo \<Rightarrow> bool) \<Rightarrow> bool"
- where "adm P \<longleftrightarrow> (\<forall>Y. chain Y \<longrightarrow> (\<forall>i. P (Y i)) \<longrightarrow> P (\<Squnion>i. Y i))"
-
-lemma admI: "(\<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i)\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)) \<Longrightarrow> adm P"
- unfolding adm_def by fast
-
-lemma admD: "adm P \<Longrightarrow> chain Y \<Longrightarrow> (\<And>i. P (Y i)) \<Longrightarrow> P (\<Squnion>i. Y i)"
- unfolding adm_def by fast
-
-lemma admD2: "adm (\<lambda>x. \<not> P x) \<Longrightarrow> chain Y \<Longrightarrow> P (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. P (Y i)"
- unfolding adm_def by fast
-
-lemma triv_admI: "\<forall>x. P x \<Longrightarrow> adm P"
- by (rule admI) (erule spec)
-
-
-subsection \<open>Admissibility on chain-finite types\<close>
-
-text \<open>For chain-finite (easy) types every formula is admissible.\<close>
-
-lemma adm_chfin [simp]: "adm P"
- for P :: "'a::chfin \<Rightarrow> bool"
- by (rule admI, frule chfin, auto simp add: maxinch_is_thelub)
-
-
-subsection \<open>Admissibility of special formulae and propagation\<close>
-
-lemma adm_const [simp]: "adm (\<lambda>x. t)"
- by (rule admI, simp)
-
-lemma adm_conj [simp]: "adm (\<lambda>x. P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<and> Q x)"
- by (fast intro: admI elim: admD)
-
-lemma adm_all [simp]: "(\<And>y. adm (\<lambda>x. P x y)) \<Longrightarrow> adm (\<lambda>x. \<forall>y. P x y)"
- by (fast intro: admI elim: admD)
-
-lemma adm_ball [simp]: "(\<And>y. y \<in> A \<Longrightarrow> adm (\<lambda>x. P x y)) \<Longrightarrow> adm (\<lambda>x. \<forall>y\<in>A. P x y)"
- by (fast intro: admI elim: admD)
-
-text \<open>Admissibility for disjunction is hard to prove. It requires 2 lemmas.\<close>
-
-lemma adm_disj_lemma1:
- assumes adm: "adm P"
- assumes chain: "chain Y"
- assumes P: "\<forall>i. \<exists>j\<ge>i. P (Y j)"
- shows "P (\<Squnion>i. Y i)"
-proof -
- define f where "f i = (LEAST j. i \<le> j \<and> P (Y j))" for i
- have chain': "chain (\<lambda>i. Y (f i))"
- unfolding f_def
- apply (rule chainI)
- apply (rule chain_mono [OF chain])
- apply (rule Least_le)
- apply (rule LeastI2_ex)
- apply (simp_all add: P)
- done
- have f1: "\<And>i. i \<le> f i" and f2: "\<And>i. P (Y (f i))"
- using LeastI_ex [OF P [rule_format]] by (simp_all add: f_def)
- have lub_eq: "(\<Squnion>i. Y i) = (\<Squnion>i. Y (f i))"
- apply (rule below_antisym)
- apply (rule lub_mono [OF chain chain'])
- apply (rule chain_mono [OF chain f1])
- apply (rule lub_range_mono [OF _ chain chain'])
- apply clarsimp
- done
- show "P (\<Squnion>i. Y i)"
- unfolding lub_eq using adm chain' f2 by (rule admD)
-qed
-
-lemma adm_disj_lemma2: "\<forall>n::nat. P n \<or> Q n \<Longrightarrow> (\<forall>i. \<exists>j\<ge>i. P j) \<or> (\<forall>i. \<exists>j\<ge>i. Q j)"
- apply (erule contrapos_pp)
- apply (clarsimp, rename_tac a b)
- apply (rule_tac x="max a b" in exI)
- apply simp
- done
-
-lemma adm_disj [simp]: "adm (\<lambda>x. P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<or> Q x)"
- apply (rule admI)
- apply (erule adm_disj_lemma2 [THEN disjE])
- apply (erule (2) adm_disj_lemma1 [THEN disjI1])
- apply (erule (2) adm_disj_lemma1 [THEN disjI2])
- done
-
-lemma adm_imp [simp]: "adm (\<lambda>x. \<not> P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<longrightarrow> Q x)"
- by (subst imp_conv_disj) (rule adm_disj)
-
-lemma adm_iff [simp]: "adm (\<lambda>x. P x \<longrightarrow> Q x) \<Longrightarrow> adm (\<lambda>x. Q x \<longrightarrow> P x) \<Longrightarrow> adm (\<lambda>x. P x \<longleftrightarrow> Q x)"
- by (subst iff_conv_conj_imp) (rule adm_conj)
-
-text \<open>admissibility and continuity\<close>
-
-lemma adm_below [simp]: "cont (\<lambda>x. u x) \<Longrightarrow> cont (\<lambda>x. v x) \<Longrightarrow> adm (\<lambda>x. u x \<sqsubseteq> v x)"
- by (simp add: adm_def cont2contlubE lub_mono ch2ch_cont)
-
-lemma adm_eq [simp]: "cont (\<lambda>x. u x) \<Longrightarrow> cont (\<lambda>x. v x) \<Longrightarrow> adm (\<lambda>x. u x = v x)"
- by (simp add: po_eq_conv)
-
-lemma adm_subst: "cont (\<lambda>x. t x) \<Longrightarrow> adm P \<Longrightarrow> adm (\<lambda>x. P (t x))"
- by (simp add: adm_def cont2contlubE ch2ch_cont)
-
-lemma adm_not_below [simp]: "cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. t x \<notsqsubseteq> u)"
- by (rule admI) (simp add: cont2contlubE ch2ch_cont lub_below_iff)
-
-
-subsection \<open>Compactness\<close>
-
-definition compact :: "'a::cpo \<Rightarrow> bool"
- where "compact k = adm (\<lambda>x. k \<notsqsubseteq> x)"
-
-lemma compactI: "adm (\<lambda>x. k \<notsqsubseteq> x) \<Longrightarrow> compact k"
- unfolding compact_def .
-
-lemma compactD: "compact k \<Longrightarrow> adm (\<lambda>x. k \<notsqsubseteq> x)"
- unfolding compact_def .
-
-lemma compactI2: "(\<And>Y. \<lbrakk>chain Y; x \<sqsubseteq> (\<Squnion>i. Y i)\<rbrakk> \<Longrightarrow> \<exists>i. x \<sqsubseteq> Y i) \<Longrightarrow> compact x"
- unfolding compact_def adm_def by fast
-
-lemma compactD2: "compact x \<Longrightarrow> chain Y \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. x \<sqsubseteq> Y i"
- unfolding compact_def adm_def by fast
-
-lemma compact_below_lub_iff: "compact x \<Longrightarrow> chain Y \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. Y i) \<longleftrightarrow> (\<exists>i. x \<sqsubseteq> Y i)"
- by (fast intro: compactD2 elim: below_lub)
-
-lemma compact_chfin [simp]: "compact x"
- for x :: "'a::chfin"
- by (rule compactI [OF adm_chfin])
-
-lemma compact_imp_max_in_chain: "chain Y \<Longrightarrow> compact (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. max_in_chain i Y"
- apply (drule (1) compactD2, simp)
- apply (erule exE, rule_tac x=i in exI)
- apply (rule max_in_chainI)
- apply (rule below_antisym)
- apply (erule (1) chain_mono)
- apply (erule (1) below_trans [OF is_ub_thelub])
- done
-
-text \<open>admissibility and compactness\<close>
-
-lemma adm_compact_not_below [simp]:
- "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. k \<notsqsubseteq> t x)"
- unfolding compact_def by (rule adm_subst)
-
-lemma adm_neq_compact [simp]: "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. t x \<noteq> k)"
- by (simp add: po_eq_conv)
-
-lemma adm_compact_neq [simp]: "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. k \<noteq> t x)"
- by (simp add: po_eq_conv)
-
-lemma compact_bottom [simp, intro]: "compact \<bottom>"
- by (rule compactI) simp
-
-text \<open>Any upward-closed predicate is admissible.\<close>
-
-lemma adm_upward:
- assumes P: "\<And>x y. \<lbrakk>P x; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> P y"
- shows "adm P"
- by (rule admI, drule spec, erule P, erule is_ub_thelub)
-
-lemmas adm_lemmas =
- adm_const adm_conj adm_all adm_ball adm_disj adm_imp adm_iff
- adm_below adm_eq adm_not_below
- adm_compact_not_below adm_compact_neq adm_neq_compact
-
-end