author | ballarin |
Wed, 24 Aug 2005 12:07:00 +0200 | |
changeset 17139 | 165c97f9bb63 |
parent 16738 | b70bac29b11d |
child 17586 | df8b2f0e462e |
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$ |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
3 |
Author: Franz Regensburger |
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 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
6 |
header {* Admissibility *} |
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 |
16079
757e1c4a8081
moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents:
16070
diff
changeset
|
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 |
|
16565 | 16 |
constdefs |
17 |
adm :: "('a::cpo \<Rightarrow> bool) \<Rightarrow> bool" |
|
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
18 |
"adm P \<equiv> \<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:
16565
diff
changeset
|
21 |
"(\<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i)\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)) \<Longrightarrow> adm P" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
22 |
apply (unfold adm_def) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
23 |
apply blast |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
24 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
25 |
|
16565 | 26 |
lemma triv_admI: "\<forall>x. P x \<Longrightarrow> adm P" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
27 |
apply (rule admI) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
28 |
apply (erule spec) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
29 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
30 |
|
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
31 |
lemma admD: "\<lbrakk>adm P; chain Y; \<forall>i. P (Y i)\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
32 |
apply (unfold adm_def) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
33 |
apply blast |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
34 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
35 |
|
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
36 |
text {* improved admissibility introduction *} |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
37 |
|
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
38 |
lemma admI2: |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
39 |
"(\<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:
16565
diff
changeset
|
40 |
\<Longrightarrow> P (\<Squnion>i. Y i)) \<Longrightarrow> adm P" |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
41 |
apply (rule admI) |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
42 |
apply (erule (1) increasing_chain_adm_lemma) |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
43 |
apply fast |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
44 |
done |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
45 |
|
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
46 |
subsection {* Admissibility on chain-finite types *} |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
47 |
|
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
48 |
text {* for chain-finite (easy) types every formula is admissible *} |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
49 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
50 |
lemma adm_max_in_chain: |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
51 |
"\<forall>Y. chain (Y::nat \<Rightarrow> 'a) \<longrightarrow> (\<exists>n. max_in_chain n Y) |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
52 |
\<Longrightarrow> adm (P::'a \<Rightarrow> bool)" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
53 |
apply (unfold adm_def) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
54 |
apply (intro strip) |
16565 | 55 |
apply (drule spec) |
56 |
apply (drule mp) |
|
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
57 |
apply assumption |
16565 | 58 |
apply (erule exE) |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
59 |
apply (simp add: maxinch_is_thelub) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
60 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
61 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
62 |
lemmas adm_chfin = chfin [THEN adm_max_in_chain, standard] |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
63 |
|
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
64 |
subsection {* Admissibility of special formulae and propagation *} |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
65 |
|
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
66 |
lemma adm_less: "\<lbrakk>cont u; cont v\<rbrakk> \<Longrightarrow> adm (\<lambda>x. u x \<sqsubseteq> v x)" |
16565 | 67 |
apply (rule admI) |
16738 | 68 |
apply (simp add: cont2contlubE) |
16565 | 69 |
apply (rule lub_mono) |
16738 | 70 |
apply (erule (1) ch2ch_cont) |
71 |
apply (erule (1) ch2ch_cont) |
|
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
72 |
apply assumption |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
73 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
74 |
|
16565 | 75 |
lemma adm_conj: "\<lbrakk>adm P; adm Q\<rbrakk> \<Longrightarrow> adm (\<lambda>x. P x \<and> Q x)" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
76 |
by (fast elim: admD intro: admI) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
77 |
|
16565 | 78 |
lemma adm_not_free: "adm (\<lambda>x. t)" |
79 |
by (rule admI, simp) |
|
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
80 |
|
16565 | 81 |
lemma adm_not_less: "cont t \<Longrightarrow> adm (\<lambda>x. \<not> t x \<sqsubseteq> u)" |
82 |
apply (rule admI) |
|
83 |
apply (drule_tac x=0 in spec) |
|
84 |
apply (erule contrapos_nn) |
|
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
85 |
apply (rule trans_less) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
86 |
prefer 2 apply (assumption) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
87 |
apply (erule cont2mono [THEN monofun_fun_arg]) |
16565 | 88 |
apply (erule is_ub_thelub) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
89 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
90 |
|
16565 | 91 |
lemma adm_all: "\<forall>y. adm (P y) \<Longrightarrow> adm (\<lambda>x. \<forall>y. P y x)" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
92 |
by (fast intro: admI elim: admD) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
93 |
|
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
94 |
lemmas adm_all2 = adm_all [rule_format] |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
95 |
|
16565 | 96 |
lemma adm_subst: "\<lbrakk>cont t; adm P\<rbrakk> \<Longrightarrow> adm (\<lambda>x. P (t x))" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
97 |
apply (rule admI) |
16738 | 98 |
apply (simp add: cont2contlubE) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
99 |
apply (erule admD) |
16738 | 100 |
apply (erule (1) ch2ch_cont) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
101 |
apply assumption |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
102 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
103 |
|
16565 | 104 |
lemma adm_UU_not_less: "adm (\<lambda>x. \<not> \<bottom> \<sqsubseteq> t x)" |
105 |
by (simp add: adm_not_free) |
|
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
106 |
|
16565 | 107 |
lemma adm_not_UU: "cont t \<Longrightarrow> adm (\<lambda>x. \<not> t x = \<bottom>)" |
108 |
by (simp add: eq_UU_iff adm_not_less) |
|
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
109 |
|
16565 | 110 |
lemma adm_eq: "\<lbrakk>cont u; cont v\<rbrakk> \<Longrightarrow> adm (\<lambda>x. u x = v x)" |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
111 |
by (simp add: po_eq_conv adm_conj adm_less) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
112 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
113 |
text {* admissibility for disjunction is hard to prove. It takes 7 Lemmas *} |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
114 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
115 |
lemma adm_disj_lemma1: |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
116 |
"\<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)" |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
117 |
apply (erule contrapos_pp) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
118 |
apply clarsimp |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
119 |
apply (rule exI) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
120 |
apply (rule conjI) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
121 |
apply (drule spec, erule mp) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
122 |
apply (rule le_maxI1) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
123 |
apply (drule spec, erule mp) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
124 |
apply (rule le_maxI2) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
125 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
126 |
|
16565 | 127 |
lemma adm_disj_lemma2: |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
128 |
"\<lbrakk>adm P; \<exists>X. chain X \<and> (\<forall>n. P (X n)) \<and> (\<Squnion>i. Y i) = (\<Squnion>i. X i)\<rbrakk> |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
129 |
\<Longrightarrow> P (\<Squnion>i. Y i)" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
130 |
by (force elim: admD) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
131 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
132 |
lemma adm_disj_lemma3: |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
133 |
"\<lbrakk>chain (Y::nat \<Rightarrow> 'a::cpo); \<forall>i. \<exists>j\<ge>i. P (Y j)\<rbrakk> |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
134 |
\<Longrightarrow> chain (\<lambda>m. Y (LEAST j. m \<le> j \<and> P (Y j)))" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
135 |
apply (rule chainI) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
136 |
apply (erule chain_mono3) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
137 |
apply (rule Least_le) |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
138 |
apply (drule_tac x="Suc i" in spec) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
139 |
apply (rule conjI) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
140 |
apply (rule Suc_leD) |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
141 |
apply (erule LeastI_ex [THEN conjunct1]) |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
142 |
apply (erule LeastI_ex [THEN conjunct2]) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
143 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
144 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
145 |
lemma adm_disj_lemma4: |
16565 | 146 |
"\<lbrakk>\<forall>i. \<exists>j\<ge>i. P (Y j)\<rbrakk> \<Longrightarrow> \<forall>m. P (Y (LEAST j::nat. m \<le> j \<and> P (Y j)))" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
147 |
apply (rule allI) |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
148 |
apply (drule_tac x=m in spec) |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
149 |
apply (erule LeastI_ex [THEN conjunct2]) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
150 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
151 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
152 |
lemma adm_disj_lemma5: |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
153 |
"\<lbrakk>chain (Y::nat \<Rightarrow> 'a::cpo); \<forall>i. \<exists>j\<ge>i. P (Y j)\<rbrakk> \<Longrightarrow> |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
154 |
(\<Squnion>m. Y m) = (\<Squnion>m. Y (LEAST j. m \<le> j \<and> P (Y j)))" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
155 |
apply (rule antisym_less) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
156 |
apply (rule lub_mono) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
157 |
apply assumption |
16565 | 158 |
apply (erule (1) adm_disj_lemma3) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
159 |
apply (rule allI) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
160 |
apply (erule chain_mono3) |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
161 |
apply (drule_tac x=k in spec) |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
162 |
apply (erule LeastI_ex [THEN conjunct1]) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
163 |
apply (rule lub_mono3) |
16565 | 164 |
apply (erule (1) adm_disj_lemma3) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
165 |
apply assumption |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
166 |
apply (rule allI) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
167 |
apply (rule exI) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
168 |
apply (rule refl_less) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
169 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
170 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
171 |
lemma adm_disj_lemma6: |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
172 |
"\<lbrakk>chain (Y::nat \<Rightarrow> 'a::cpo); \<forall>i. \<exists>j\<ge>i. P(Y j)\<rbrakk> \<Longrightarrow> |
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
173 |
\<exists>X. chain X \<and> (\<forall>n. P (X n)) \<and> (\<Squnion>i. Y i) = (\<Squnion>i. X i)" |
16565 | 174 |
apply (rule_tac x = "\<lambda>m. Y (LEAST j. m \<le> j \<and> P (Y j))" in exI) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
175 |
apply (fast intro!: adm_disj_lemma3 adm_disj_lemma4 adm_disj_lemma5) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
176 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
177 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
178 |
lemma adm_disj_lemma7: |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
179 |
"\<lbrakk>adm P; chain Y; \<forall>i. \<exists>j\<ge>i. P (Y j)\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)" |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
180 |
apply (erule adm_disj_lemma2) |
16565 | 181 |
apply (erule (1) adm_disj_lemma6) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
182 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
183 |
|
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
184 |
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
|
185 |
apply (rule admI) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
186 |
apply (erule adm_disj_lemma1 [THEN disjE]) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
187 |
apply (rule disjI1) |
16565 | 188 |
apply (erule (2) adm_disj_lemma7) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
189 |
apply (rule disjI2) |
16565 | 190 |
apply (erule (2) adm_disj_lemma7) |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
191 |
done |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
192 |
|
16565 | 193 |
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
|
194 |
by (subst imp_conv_disj, rule adm_disj) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
195 |
|
16565 | 196 |
lemma adm_iff: |
197 |
"\<lbrakk>adm (\<lambda>x. P x \<longrightarrow> Q x); adm (\<lambda>x. Q x \<longrightarrow> P x)\<rbrakk> |
|
198 |
\<Longrightarrow> adm (\<lambda>x. P x = Q x)" |
|
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
199 |
by (subst iff_conv_conj_imp, rule adm_conj) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
200 |
|
16565 | 201 |
lemma adm_not_conj: |
202 |
"\<lbrakk>adm (\<lambda>x. \<not> P x); adm (\<lambda>x. \<not> Q x)\<rbrakk> \<Longrightarrow> adm (\<lambda>x. \<not> (P x \<and> Q x))" |
|
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
203 |
by (subst de_Morgan_conj, rule adm_disj) |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
204 |
|
16565 | 205 |
lemmas adm_lemmas = |
16623
f3fcfa388ecb
cleaned up; added adm_less to adm_lemmas; added subsection headings
huffman
parents:
16565
diff
changeset
|
206 |
adm_less adm_conj adm_not_free adm_imp adm_disj adm_eq adm_not_UU |
16565 | 207 |
adm_UU_not_less adm_all2 adm_not_less adm_not_conj adm_iff |
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
208 |
|
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
209 |
declare adm_lemmas [simp] |
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
210 |
|
16062
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
211 |
(* legacy ML bindings *) |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
212 |
ML |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
213 |
{* |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
214 |
val adm_def = thm "adm_def"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
215 |
val admI = thm "admI"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
216 |
val triv_admI = thm "triv_admI"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
217 |
val admD = thm "admD"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
218 |
val adm_max_in_chain = thm "adm_max_in_chain"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
219 |
val adm_chfin = thm "adm_chfin"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
220 |
val admI2 = thm "admI2"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
221 |
val adm_less = thm "adm_less"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
222 |
val adm_conj = thm "adm_conj"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
223 |
val adm_not_free = thm "adm_not_free"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
224 |
val adm_not_less = thm "adm_not_less"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
225 |
val adm_all = thm "adm_all"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
226 |
val adm_all2 = thm "adm_all2"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
227 |
val adm_subst = thm "adm_subst"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
228 |
val adm_UU_not_less = thm "adm_UU_not_less"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
229 |
val adm_not_UU = thm "adm_not_UU"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
230 |
val adm_eq = thm "adm_eq"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
231 |
val adm_disj_lemma1 = thm "adm_disj_lemma1"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
232 |
val adm_disj_lemma2 = thm "adm_disj_lemma2"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
233 |
val adm_disj_lemma3 = thm "adm_disj_lemma3"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
234 |
val adm_disj_lemma4 = thm "adm_disj_lemma4"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
235 |
val adm_disj_lemma5 = thm "adm_disj_lemma5"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
236 |
val adm_disj_lemma6 = thm "adm_disj_lemma6"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
237 |
val adm_disj_lemma7 = thm "adm_disj_lemma7"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
238 |
val adm_disj = thm "adm_disj"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
239 |
val adm_imp = thm "adm_imp"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
240 |
val adm_iff = thm "adm_iff"; |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
241 |
val adm_not_conj = thm "adm_not_conj"; |
16565 | 242 |
val adm_lemmas = thms "adm_lemmas"; |
16062
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
243 |
*} |
f8110bd9957f
cannot have files named adm.ML and Adm.ML on Macs, so deleted one and renamed the other
paulson
parents:
16056
diff
changeset
|
244 |
|
16056
32c3b7188c28
Moved admissibility definitions and lemmas to a separate theory
huffman
parents:
diff
changeset
|
245 |
end |