| author | wenzelm | 
| Wed, 01 Oct 2008 12:00:00 +0200 | |
| changeset 28441 | 9b0daffc4054 | 
| parent 27419 | ff2a2b8fcd09 | 
| child 29138 | 661a8db7e647 | 
| permissions | -rw-r--r-- | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 1 | (* Title: HOLCF/Adm.thy | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 2 | ID: $Id$ | 
| 25895 
0eaadfa8889e
converted adm_all and adm_ball to rule_format; cleaned up
 huffman parents: 
25880diff
changeset | 3 | Author: Franz Regensburger and Brian Huffman | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 4 | *) | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 5 | |
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 6 | header {* Admissibility and compactness *}
 | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 7 | |
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 8 | theory Adm | 
| 27181 | 9 | imports Cont | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 10 | begin | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 11 | |
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 12 | defaultsort cpo | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 13 | |
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 14 | subsection {* Definitions *}
 | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 15 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19440diff
changeset | 16 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19440diff
changeset | 17 |   adm :: "('a::cpo \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19440diff
changeset | 18 | "adm P = (\<forall>Y. chain Y \<longrightarrow> (\<forall>i. P (Y i)) \<longrightarrow> P (\<Squnion>i. Y i))" | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 19 | |
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 20 | lemma admI: | 
| 16623 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 21 | "(\<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i)\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)) \<Longrightarrow> adm P" | 
| 25895 
0eaadfa8889e
converted adm_all and adm_ball to rule_format; cleaned up
 huffman parents: 
25880diff
changeset | 22 | unfolding adm_def by fast | 
| 
0eaadfa8889e
converted adm_all and adm_ball to rule_format; cleaned up
 huffman parents: 
25880diff
changeset | 23 | |
| 25925 | 24 | lemma admD: "\<lbrakk>adm P; chain Y; \<And>i. P (Y i)\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)" | 
| 25895 
0eaadfa8889e
converted adm_all and adm_ball to rule_format; cleaned up
 huffman parents: 
25880diff
changeset | 25 | unfolding adm_def by fast | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 26 | |
| 27181 | 27 | lemma admD2: "\<lbrakk>adm (\<lambda>x. \<not> P x); chain Y; P (\<Squnion>i. Y i)\<rbrakk> \<Longrightarrow> \<exists>i. P (Y i)" | 
| 28 | unfolding adm_def by fast | |
| 29 | ||
| 16565 | 30 | lemma triv_admI: "\<forall>x. P x \<Longrightarrow> adm P" | 
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 31 | by (rule admI, erule spec) | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 32 | |
| 16623 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 33 | text {* improved admissibility introduction *}
 | 
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 34 | |
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 35 | lemma admI2: | 
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 36 | "(\<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i); \<forall>i. \<exists>j>i. Y i \<noteq> Y j \<and> Y i \<sqsubseteq> Y j\<rbrakk> | 
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 37 | \<Longrightarrow> P (\<Squnion>i. Y i)) \<Longrightarrow> adm P" | 
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 38 | apply (rule admI) | 
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 39 | apply (erule (1) increasing_chain_adm_lemma) | 
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 40 | apply fast | 
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 41 | done | 
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 42 | |
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 43 | subsection {* Admissibility on chain-finite types *}
 | 
| 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 44 | |
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 45 | text {* for chain-finite (easy) types every formula is admissible *}
 | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 46 | |
| 25921 | 47 | lemma adm_chfin: "adm (P::'a::chfin \<Rightarrow> bool)" | 
| 48 | by (rule admI, frule chfin, auto simp add: maxinch_is_thelub) | |
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 49 | |
| 16623 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 50 | subsection {* Admissibility of special formulae and propagation *}
 | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 51 | |
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 52 | lemma adm_not_free: "adm (\<lambda>x. t)" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 53 | by (rule admI, simp) | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 54 | |
| 16565 | 55 | lemma adm_conj: "\<lbrakk>adm P; adm Q\<rbrakk> \<Longrightarrow> adm (\<lambda>x. P x \<and> Q x)" | 
| 25925 | 56 | by (fast intro: admI elim: admD) | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 57 | |
| 27290 | 58 | lemma adm_all: "(\<And>y. adm (\<lambda>x. P x y)) \<Longrightarrow> adm (\<lambda>x. \<forall>y. P x y)" | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 59 | by (fast intro: admI elim: admD) | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 60 | |
| 27290 | 61 | lemma adm_ball: "(\<And>y. y \<in> A \<Longrightarrow> adm (\<lambda>x. P x y)) \<Longrightarrow> adm (\<lambda>x. \<forall>y\<in>A. P x y)" | 
| 17586 | 62 | by (fast intro: admI elim: admD) | 
| 63 | ||
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 64 | text {* Admissibility for disjunction is hard to prove. It takes 5 Lemmas *}
 | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 65 | |
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 66 | lemma adm_disj_lemma1: | 
| 16623 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 67 | "\<lbrakk>chain (Y::nat \<Rightarrow> 'a::cpo); \<forall>i. \<exists>j\<ge>i. P (Y j)\<rbrakk> | 
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 68 | \<Longrightarrow> chain (\<lambda>i. Y (LEAST j. i \<le> j \<and> P (Y j)))" | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 69 | apply (rule chainI) | 
| 25922 
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
 huffman parents: 
25921diff
changeset | 70 | apply (erule chain_mono) | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 71 | apply (rule Least_le) | 
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 72 | apply (rule LeastI2_ex) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 73 | apply simp_all | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 74 | done | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 75 | |
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 76 | lemmas adm_disj_lemma2 = LeastI_ex [of "\<lambda>j. i \<le> j \<and> P (Y j)", standard] | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 77 | |
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 78 | lemma adm_disj_lemma3: | 
| 16623 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 79 | "\<lbrakk>chain (Y::nat \<Rightarrow> 'a::cpo); \<forall>i. \<exists>j\<ge>i. P (Y j)\<rbrakk> \<Longrightarrow> | 
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 80 | (\<Squnion>i. Y i) = (\<Squnion>i. Y (LEAST j. i \<le> j \<and> P (Y j)))" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 81 | apply (frule (1) adm_disj_lemma1) | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 82 | apply (rule antisym_less) | 
| 25923 | 83 | apply (rule lub_mono, assumption+) | 
| 25922 
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
 huffman parents: 
25921diff
changeset | 84 | apply (erule chain_mono) | 
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 85 | apply (simp add: adm_disj_lemma2) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 86 | apply (rule lub_range_mono, fast, assumption+) | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 87 | done | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 88 | |
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 89 | lemma adm_disj_lemma4: | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 90 | "\<lbrakk>adm P; chain Y; \<forall>i. \<exists>j\<ge>i. P (Y j)\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 91 | apply (subst adm_disj_lemma3, assumption+) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 92 | apply (erule admD) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 93 | apply (simp add: adm_disj_lemma1) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 94 | apply (simp add: adm_disj_lemma2) | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 95 | done | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 96 | |
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 97 | lemma adm_disj_lemma5: | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 98 | "\<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)" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 99 | apply (erule contrapos_pp) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 100 | apply (clarsimp, rename_tac a b) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 101 | apply (rule_tac x="max a b" in exI) | 
| 25895 
0eaadfa8889e
converted adm_all and adm_ball to rule_format; cleaned up
 huffman parents: 
25880diff
changeset | 102 | apply simp | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 103 | done | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 104 | |
| 16623 
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
 huffman parents: 
16565diff
changeset | 105 | lemma adm_disj: "\<lbrakk>adm P; adm Q\<rbrakk> \<Longrightarrow> adm (\<lambda>x. P x \<or> Q x)" | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 106 | apply (rule admI) | 
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 107 | apply (erule adm_disj_lemma5 [THEN disjE]) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 108 | apply (erule (2) adm_disj_lemma4 [THEN disjI1]) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 109 | apply (erule (2) adm_disj_lemma4 [THEN disjI2]) | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 110 | done | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 111 | |
| 16565 | 112 | lemma adm_imp: "\<lbrakk>adm (\<lambda>x. \<not> P x); adm Q\<rbrakk> \<Longrightarrow> adm (\<lambda>x. P x \<longrightarrow> Q x)" | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 113 | by (subst imp_conv_disj, rule adm_disj) | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 114 | |
| 16565 | 115 | lemma adm_iff: | 
| 116 | "\<lbrakk>adm (\<lambda>x. P x \<longrightarrow> Q x); adm (\<lambda>x. Q x \<longrightarrow> P x)\<rbrakk> | |
| 117 | \<Longrightarrow> adm (\<lambda>x. P x = Q x)" | |
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 118 | by (subst iff_conv_conj_imp, rule adm_conj) | 
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 119 | |
| 16565 | 120 | lemma adm_not_conj: | 
| 121 | "\<lbrakk>adm (\<lambda>x. \<not> P x); adm (\<lambda>x. \<not> Q x)\<rbrakk> \<Longrightarrow> adm (\<lambda>x. \<not> (P x \<and> Q x))" | |
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 122 | by (simp add: adm_imp) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 123 | |
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 124 | text {* admissibility and continuity *}
 | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 125 | |
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 126 | lemma adm_less: "\<lbrakk>cont u; cont v\<rbrakk> \<Longrightarrow> adm (\<lambda>x. u x \<sqsubseteq> v x)" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 127 | apply (rule admI) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 128 | apply (simp add: cont2contlubE) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 129 | apply (rule lub_mono) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 130 | apply (erule (1) ch2ch_cont) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 131 | apply (erule (1) ch2ch_cont) | 
| 25923 | 132 | apply (erule spec) | 
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 133 | done | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 134 | |
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 135 | lemma adm_eq: "\<lbrakk>cont u; cont v\<rbrakk> \<Longrightarrow> adm (\<lambda>x. u x = v x)" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 136 | by (simp add: po_eq_conv adm_conj adm_less) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 137 | |
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 138 | lemma adm_subst: "\<lbrakk>cont t; adm P\<rbrakk> \<Longrightarrow> adm (\<lambda>x. P (t x))" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 139 | apply (rule admI) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 140 | apply (simp add: cont2contlubE) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 141 | apply (erule admD) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 142 | apply (erule (1) ch2ch_cont) | 
| 25925 | 143 | apply (erule spec) | 
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 144 | done | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 145 | |
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 146 | lemma adm_not_less: "cont t \<Longrightarrow> adm (\<lambda>x. \<not> t x \<sqsubseteq> u)" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 147 | apply (rule admI) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 148 | apply (drule_tac x=0 in spec) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 149 | apply (erule contrapos_nn) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 150 | apply (erule rev_trans_less) | 
| 25786 | 151 | apply (erule cont2mono [THEN monofunE]) | 
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 152 | apply (erule is_ub_thelub) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 153 | done | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 154 | |
| 25880 | 155 | subsection {* Compactness *}
 | 
| 156 | ||
| 157 | definition | |
| 158 | compact :: "'a::cpo \<Rightarrow> bool" where | |
| 159 | "compact k = adm (\<lambda>x. \<not> k \<sqsubseteq> x)" | |
| 160 | ||
| 161 | lemma compactI: "adm (\<lambda>x. \<not> k \<sqsubseteq> x) \<Longrightarrow> compact k" | |
| 162 | unfolding compact_def . | |
| 163 | ||
| 164 | lemma compactD: "compact k \<Longrightarrow> adm (\<lambda>x. \<not> k \<sqsubseteq> x)" | |
| 165 | unfolding compact_def . | |
| 166 | ||
| 167 | lemma compactI2: | |
| 27413 | 168 | "(\<And>Y. \<lbrakk>chain Y; x \<sqsubseteq> (\<Squnion>i. Y i)\<rbrakk> \<Longrightarrow> \<exists>i. x \<sqsubseteq> Y i) \<Longrightarrow> compact x" | 
| 25880 | 169 | unfolding compact_def adm_def by fast | 
| 170 | ||
| 171 | lemma compactD2: | |
| 27413 | 172 | "\<lbrakk>compact x; chain Y; x \<sqsubseteq> (\<Squnion>i. Y i)\<rbrakk> \<Longrightarrow> \<exists>i. x \<sqsubseteq> Y i" | 
| 25880 | 173 | unfolding compact_def adm_def by fast | 
| 174 | ||
| 175 | lemma compact_chfin [simp]: "compact (x::'a::chfin)" | |
| 176 | by (rule compactI [OF adm_chfin]) | |
| 177 | ||
| 178 | lemma compact_imp_max_in_chain: | |
| 179 | "\<lbrakk>chain Y; compact (\<Squnion>i. Y i)\<rbrakk> \<Longrightarrow> \<exists>i. max_in_chain i Y" | |
| 180 | apply (drule (1) compactD2, simp) | |
| 181 | apply (erule exE, rule_tac x=i in exI) | |
| 182 | apply (rule max_in_chainI) | |
| 183 | apply (rule antisym_less) | |
| 25922 
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
 huffman parents: 
25921diff
changeset | 184 | apply (erule (1) chain_mono) | 
| 25880 | 185 | apply (erule (1) trans_less [OF is_ub_thelub]) | 
| 186 | done | |
| 187 | ||
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 188 | text {* admissibility and compactness *}
 | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 189 | |
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 190 | lemma adm_compact_not_less: "\<lbrakk>compact k; cont t\<rbrakk> \<Longrightarrow> adm (\<lambda>x. \<not> k \<sqsubseteq> t x)" | 
| 25880 | 191 | unfolding compact_def by (rule adm_subst) | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 192 | |
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 193 | lemma adm_neq_compact: "\<lbrakk>compact k; cont t\<rbrakk> \<Longrightarrow> adm (\<lambda>x. t x \<noteq> k)" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 194 | by (simp add: po_eq_conv adm_imp adm_not_less adm_compact_not_less) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 195 | |
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 196 | lemma adm_compact_neq: "\<lbrakk>compact k; cont t\<rbrakk> \<Longrightarrow> adm (\<lambda>x. k \<noteq> t x)" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 197 | by (simp add: po_eq_conv adm_imp adm_not_less adm_compact_not_less) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 198 | |
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 199 | lemma compact_UU [simp, intro]: "compact \<bottom>" | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 200 | by (rule compactI, simp add: adm_not_free) | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 201 | |
| 19440 | 202 | lemma adm_not_UU: "cont t \<Longrightarrow> adm (\<lambda>x. t x \<noteq> \<bottom>)" | 
| 203 | by (simp add: adm_neq_compact) | |
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 204 | |
| 25802 | 205 | text {* Any upward-closed predicate is admissible. *}
 | 
| 206 | ||
| 207 | lemma adm_upward: | |
| 208 | assumes P: "\<And>x y. \<lbrakk>P x; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> P y" | |
| 209 | shows "adm P" | |
| 210 | by (rule admI, drule spec, erule P, erule is_ub_thelub) | |
| 211 | ||
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 212 | lemmas adm_lemmas [simp] = | 
| 25895 
0eaadfa8889e
converted adm_all and adm_ball to rule_format; cleaned up
 huffman parents: 
25880diff
changeset | 213 | adm_not_free adm_conj adm_all adm_ball adm_disj adm_imp adm_iff | 
| 17814 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 214 | adm_less adm_eq adm_not_less | 
| 
21183d6f62b8
added notion of compactness; shortened proof of adm_disj; reorganized and cleaned up
 huffman parents: 
17586diff
changeset | 215 | adm_compact_not_less adm_compact_neq adm_neq_compact adm_not_UU | 
| 16056 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 216 | |
| 
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
 huffman parents: diff
changeset | 217 | end |