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