| author | wenzelm | 
| Mon, 31 Dec 2018 20:08:32 +0100 | |
| changeset 69559 | 66c8dff9639f | 
| parent 67312 | 0d25e02759b7 | 
| permissions | -rw-r--r-- | 
| 41959 | 1  | 
(* Title: HOL/HOLCF/Adm.thy  | 
| 
25895
 
0eaadfa8889e
converted adm_all and adm_ball to rule_format; cleaned up
 
huffman 
parents: 
25880 
diff
changeset
 | 
2  | 
Author: Franz Regensburger and Brian Huffman  | 
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
4  | 
|
| 62175 | 5  | 
section \<open>Admissibility and compactness\<close>  | 
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
7  | 
theory Adm  | 
| 67312 | 8  | 
imports Cont  | 
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
9  | 
begin  | 
| 
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
10  | 
|
| 36452 | 11  | 
default_sort cpo  | 
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
12  | 
|
| 62175 | 13  | 
subsection \<open>Definitions\<close>  | 
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
14  | 
|
| 67312 | 15  | 
definition adm :: "('a::cpo \<Rightarrow> bool) \<Rightarrow> bool"
 | 
16  | 
where "adm P \<longleftrightarrow> (\<forall>Y. chain Y \<longrightarrow> (\<forall>i. P (Y i)) \<longrightarrow> P (\<Squnion>i. Y i))"  | 
|
17  | 
||
18  | 
lemma admI: "(\<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i)\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)) \<Longrightarrow> adm P"  | 
|
19  | 
unfolding adm_def by fast  | 
|
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
20  | 
|
| 67312 | 21  | 
lemma admD: "adm P \<Longrightarrow> chain Y \<Longrightarrow> (\<And>i. P (Y i)) \<Longrightarrow> P (\<Squnion>i. Y i)"  | 
22  | 
unfolding adm_def by fast  | 
|
| 
25895
 
0eaadfa8889e
converted adm_all and adm_ball to rule_format; cleaned up
 
huffman 
parents: 
25880 
diff
changeset
 | 
23  | 
|
| 67312 | 24  | 
lemma admD2: "adm (\<lambda>x. \<not> P x) \<Longrightarrow> chain Y \<Longrightarrow> P (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. P (Y i)"  | 
25  | 
unfolding adm_def by fast  | 
|
| 27181 | 26  | 
|
| 16565 | 27  | 
lemma triv_admI: "\<forall>x. P x \<Longrightarrow> adm P"  | 
| 67312 | 28  | 
by (rule admI) (erule spec)  | 
29  | 
||
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
30  | 
|
| 62175 | 31  | 
subsection \<open>Admissibility on chain-finite types\<close>  | 
| 
16623
 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 
huffman 
parents: 
16565 
diff
changeset
 | 
32  | 
|
| 62175 | 33  | 
text \<open>For chain-finite (easy) types every formula is admissible.\<close>  | 
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
34  | 
|
| 67312 | 35  | 
lemma adm_chfin [simp]: "adm P"  | 
36  | 
for P :: "'a::chfin \<Rightarrow> bool"  | 
|
37  | 
by (rule admI, frule chfin, auto simp add: maxinch_is_thelub)  | 
|
38  | 
||
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
39  | 
|
| 62175 | 40  | 
subsection \<open>Admissibility of special formulae and propagation\<close>  | 
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
41  | 
|
| 40007 | 42  | 
lemma adm_const [simp]: "adm (\<lambda>x. t)"  | 
| 67312 | 43  | 
by (rule admI, simp)  | 
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
44  | 
|
| 67312 | 45  | 
lemma adm_conj [simp]: "adm (\<lambda>x. P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<and> Q x)"  | 
46  | 
by (fast intro: admI elim: admD)  | 
|
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
47  | 
|
| 67312 | 48  | 
lemma adm_all [simp]: "(\<And>y. adm (\<lambda>x. P x y)) \<Longrightarrow> adm (\<lambda>x. \<forall>y. P x y)"  | 
49  | 
by (fast intro: admI elim: admD)  | 
|
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
50  | 
|
| 67312 | 51  | 
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)"  | 
52  | 
by (fast intro: admI elim: admD)  | 
|
| 17586 | 53  | 
|
| 62175 | 54  | 
text \<open>Admissibility for disjunction is hard to prove. It requires 2 lemmas.\<close>  | 
| 
17814
 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 
huffman 
parents: 
17586 
diff
changeset
 | 
55  | 
|
| 40007 | 56  | 
lemma adm_disj_lemma1:  | 
57  | 
assumes adm: "adm P"  | 
|
58  | 
assumes chain: "chain Y"  | 
|
59  | 
assumes P: "\<forall>i. \<exists>j\<ge>i. P (Y j)"  | 
|
60  | 
shows "P (\<Squnion>i. Y i)"  | 
|
61  | 
proof -  | 
|
| 63040 | 62  | 
define f where "f i = (LEAST j. i \<le> j \<and> P (Y j))" for i  | 
| 40007 | 63  | 
have chain': "chain (\<lambda>i. Y (f i))"  | 
64  | 
unfolding f_def  | 
|
65  | 
apply (rule chainI)  | 
|
66  | 
apply (rule chain_mono [OF chain])  | 
|
67  | 
apply (rule Least_le)  | 
|
68  | 
apply (rule LeastI2_ex)  | 
|
| 67312 | 69  | 
apply (simp_all add: P)  | 
| 40007 | 70  | 
done  | 
71  | 
have f1: "\<And>i. i \<le> f i" and f2: "\<And>i. P (Y (f i))"  | 
|
72  | 
using LeastI_ex [OF P [rule_format]] by (simp_all add: f_def)  | 
|
73  | 
have lub_eq: "(\<Squnion>i. Y i) = (\<Squnion>i. Y (f i))"  | 
|
74  | 
apply (rule below_antisym)  | 
|
| 67312 | 75  | 
apply (rule lub_mono [OF chain chain'])  | 
76  | 
apply (rule chain_mono [OF chain f1])  | 
|
| 40007 | 77  | 
apply (rule lub_range_mono [OF _ chain chain'])  | 
78  | 
apply clarsimp  | 
|
79  | 
done  | 
|
80  | 
show "P (\<Squnion>i. Y i)"  | 
|
81  | 
unfolding lub_eq using adm chain' f2 by (rule admD)  | 
|
82  | 
qed  | 
|
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
83  | 
|
| 67312 | 84  | 
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)"  | 
85  | 
apply (erule contrapos_pp)  | 
|
86  | 
apply (clarsimp, rename_tac a b)  | 
|
87  | 
apply (rule_tac x="max a b" in exI)  | 
|
88  | 
apply simp  | 
|
89  | 
done  | 
|
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
90  | 
|
| 67312 | 91  | 
lemma adm_disj [simp]: "adm (\<lambda>x. P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<or> Q x)"  | 
92  | 
apply (rule admI)  | 
|
93  | 
apply (erule adm_disj_lemma2 [THEN disjE])  | 
|
94  | 
apply (erule (2) adm_disj_lemma1 [THEN disjI1])  | 
|
95  | 
apply (erule (2) adm_disj_lemma1 [THEN disjI2])  | 
|
96  | 
done  | 
|
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
97  | 
|
| 67312 | 98  | 
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)"  | 
99  | 
by (subst imp_conv_disj) (rule adm_disj)  | 
|
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
100  | 
|
| 67312 | 101  | 
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)"  | 
102  | 
by (subst iff_conv_conj_imp) (rule adm_conj)  | 
|
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
103  | 
|
| 62175 | 104  | 
text \<open>admissibility and continuity\<close>  | 
| 
17814
 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 
huffman 
parents: 
17586 
diff
changeset
 | 
105  | 
|
| 67312 | 106  | 
lemma adm_below [simp]: "cont (\<lambda>x. u x) \<Longrightarrow> cont (\<lambda>x. v x) \<Longrightarrow> adm (\<lambda>x. u x \<sqsubseteq> v x)"  | 
107  | 
by (simp add: adm_def cont2contlubE lub_mono ch2ch_cont)  | 
|
| 
17814
 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 
huffman 
parents: 
17586 
diff
changeset
 | 
108  | 
|
| 67312 | 109  | 
lemma adm_eq [simp]: "cont (\<lambda>x. u x) \<Longrightarrow> cont (\<lambda>x. v x) \<Longrightarrow> adm (\<lambda>x. u x = v x)"  | 
110  | 
by (simp add: po_eq_conv)  | 
|
| 
17814
 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 
huffman 
parents: 
17586 
diff
changeset
 | 
111  | 
|
| 67312 | 112  | 
lemma adm_subst: "cont (\<lambda>x. t x) \<Longrightarrow> adm P \<Longrightarrow> adm (\<lambda>x. P (t x))"  | 
113  | 
by (simp add: adm_def cont2contlubE ch2ch_cont)  | 
|
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
114  | 
|
| 41182 | 115  | 
lemma adm_not_below [simp]: "cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. t x \<notsqsubseteq> u)"  | 
| 67312 | 116  | 
by (rule admI) (simp add: cont2contlubE ch2ch_cont lub_below_iff)  | 
117  | 
||
| 
17814
 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 
huffman 
parents: 
17586 
diff
changeset
 | 
118  | 
|
| 62175 | 119  | 
subsection \<open>Compactness\<close>  | 
| 25880 | 120  | 
|
| 67312 | 121  | 
definition compact :: "'a::cpo \<Rightarrow> bool"  | 
122  | 
where "compact k = adm (\<lambda>x. k \<notsqsubseteq> x)"  | 
|
| 25880 | 123  | 
|
| 41182 | 124  | 
lemma compactI: "adm (\<lambda>x. k \<notsqsubseteq> x) \<Longrightarrow> compact k"  | 
| 67312 | 125  | 
unfolding compact_def .  | 
| 25880 | 126  | 
|
| 41182 | 127  | 
lemma compactD: "compact k \<Longrightarrow> adm (\<lambda>x. k \<notsqsubseteq> x)"  | 
| 67312 | 128  | 
unfolding compact_def .  | 
129  | 
||
130  | 
lemma compactI2: "(\<And>Y. \<lbrakk>chain Y; x \<sqsubseteq> (\<Squnion>i. Y i)\<rbrakk> \<Longrightarrow> \<exists>i. x \<sqsubseteq> Y i) \<Longrightarrow> compact x"  | 
|
131  | 
unfolding compact_def adm_def by fast  | 
|
| 25880 | 132  | 
|
| 67312 | 133  | 
lemma compactD2: "compact x \<Longrightarrow> chain Y \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. x \<sqsubseteq> Y i"  | 
134  | 
unfolding compact_def adm_def by fast  | 
|
| 25880 | 135  | 
|
| 67312 | 136  | 
lemma compact_below_lub_iff: "compact x \<Longrightarrow> chain Y \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. Y i) \<longleftrightarrow> (\<exists>i. x \<sqsubseteq> Y i)"  | 
137  | 
by (fast intro: compactD2 elim: below_lub)  | 
|
| 25880 | 138  | 
|
| 67312 | 139  | 
lemma compact_chfin [simp]: "compact x"  | 
140  | 
for x :: "'a::chfin"  | 
|
141  | 
by (rule compactI [OF adm_chfin])  | 
|
| 25880 | 142  | 
|
| 67312 | 143  | 
lemma compact_imp_max_in_chain: "chain Y \<Longrightarrow> compact (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. max_in_chain i Y"  | 
144  | 
apply (drule (1) compactD2, simp)  | 
|
145  | 
apply (erule exE, rule_tac x=i in exI)  | 
|
146  | 
apply (rule max_in_chainI)  | 
|
147  | 
apply (rule below_antisym)  | 
|
148  | 
apply (erule (1) chain_mono)  | 
|
149  | 
apply (erule (1) below_trans [OF is_ub_thelub])  | 
|
150  | 
done  | 
|
| 25880 | 151  | 
|
| 62175 | 152  | 
text \<open>admissibility and compactness\<close>  | 
| 
17814
 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 
huffman 
parents: 
17586 
diff
changeset
 | 
153  | 
|
| 40007 | 154  | 
lemma adm_compact_not_below [simp]:  | 
| 67312 | 155  | 
"compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. k \<notsqsubseteq> t x)"  | 
156  | 
unfolding compact_def by (rule adm_subst)  | 
|
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
157  | 
|
| 67312 | 158  | 
lemma adm_neq_compact [simp]: "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. t x \<noteq> k)"  | 
159  | 
by (simp add: po_eq_conv)  | 
|
| 
17814
 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 
huffman 
parents: 
17586 
diff
changeset
 | 
160  | 
|
| 67312 | 161  | 
lemma adm_compact_neq [simp]: "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. k \<noteq> t x)"  | 
162  | 
by (simp add: po_eq_conv)  | 
|
| 
17814
 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 
huffman 
parents: 
17586 
diff
changeset
 | 
163  | 
|
| 
41430
 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 
huffman 
parents: 
41182 
diff
changeset
 | 
164  | 
lemma compact_bottom [simp, intro]: "compact \<bottom>"  | 
| 67312 | 165  | 
by (rule compactI) simp  | 
| 
17814
 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 
huffman 
parents: 
17586 
diff
changeset
 | 
166  | 
|
| 62175 | 167  | 
text \<open>Any upward-closed predicate is admissible.\<close>  | 
| 25802 | 168  | 
|
169  | 
lemma adm_upward:  | 
|
170  | 
assumes P: "\<And>x y. \<lbrakk>P x; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> P y"  | 
|
171  | 
shows "adm P"  | 
|
| 67312 | 172  | 
by (rule admI, drule spec, erule P, erule is_ub_thelub)  | 
| 25802 | 173  | 
|
| 40007 | 174  | 
lemmas adm_lemmas =  | 
175  | 
adm_const adm_conj adm_all adm_ball adm_disj adm_imp adm_iff  | 
|
| 
31076
 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 
huffman 
parents: 
29138 
diff
changeset
 | 
176  | 
adm_below adm_eq adm_not_below  | 
| 40007 | 177  | 
adm_compact_not_below adm_compact_neq adm_neq_compact  | 
| 
16056
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
178  | 
|
| 
 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 
huffman 
parents:  
diff
changeset
 | 
179  | 
end  |