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