src/HOL/FixedPoint.thy
author haftmann
Mon, 13 Nov 2006 15:43:06 +0100
changeset 21329 7338206d75f1
parent 21326 c33cdc5a6c7c
child 21404 eb85850d3eb7
permissions -rw-r--r--
introduces preorders
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     1
(*  Title:      HOL/FixedPoint.thy
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     2
    ID:         $Id$
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
     4
    Author:     Stefan Berghofer, TU Muenchen
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     5
    Copyright   1992  University of Cambridge
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     6
*)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     7
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     8
header{* Fixed Points and the Knaster-Tarski Theorem*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     9
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    10
theory FixedPoint
21326
c33cdc5a6c7c added LOrder dependency
haftmann
parents: 21316
diff changeset
    11
imports Product_Type LOrder
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    12
begin
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    13
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    14
subsection {* Complete lattices *}
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    15
(*FIXME Meet \<rightarrow> Inf *)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    16
consts
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    17
  Meet :: "'a::order set \<Rightarrow> 'a"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    18
  Sup :: "'a::order set \<Rightarrow> 'a"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    19
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    20
defs Sup_def: "Sup A == Meet {b. \<forall>a \<in> A. a <= b}"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    21
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    22
definition
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    23
 SUP :: "('a \<Rightarrow> 'b::order) \<Rightarrow> 'b" (binder "SUP " 10)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    24
"SUP x. f x == Sup (f ` UNIV)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    25
(*
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    26
abbreviation
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    27
 bot :: "'a::order"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    28
"bot == Sup {}"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    29
*)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    30
axclass comp_lat < order
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    31
  Meet_lower: "x \<in> A \<Longrightarrow> Meet A <= x"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    32
  Meet_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z <= x) \<Longrightarrow> z <= Meet A"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    33
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    34
theorem Sup_upper: "(x::'a::comp_lat) \<in> A \<Longrightarrow> x <= Sup A"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    35
  by (auto simp: Sup_def intro: Meet_greatest)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    36
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    37
theorem Sup_least: "(\<And>x::'a::comp_lat. x \<in> A \<Longrightarrow> x <= z) \<Longrightarrow> Sup A <= z"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    38
  by (auto simp: Sup_def intro: Meet_lower)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    39
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    40
text {* A complete lattice is a lattice *}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    41
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    42
lemma is_meet_Meet: "is_meet (\<lambda>(x::'a::comp_lat) y. Meet {x, y})"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    43
  by (auto simp: is_meet_def intro: Meet_lower Meet_greatest)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    44
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    45
lemma is_join_Sup: "is_join (\<lambda>(x::'a::comp_lat) y. Sup {x, y})"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    46
  by (auto simp: is_join_def intro: Sup_upper Sup_least)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    47
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    48
instance comp_lat < lorder
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    49
proof
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    50
  from is_meet_Meet show "\<exists>m::'a\<Rightarrow>'a\<Rightarrow>'a. is_meet m" by iprover
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    51
  from is_join_Sup show "\<exists>j::'a\<Rightarrow>'a\<Rightarrow>'a. is_join j" by iprover
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    52
qed
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    53
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    54
subsubsection {* Properties *}
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    55
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    56
lemma mono_join: "mono f \<Longrightarrow> join (f A) (f B) <= f (join A B)"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    57
  by (auto simp add: mono_def)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    58
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    59
lemma mono_meet: "mono f \<Longrightarrow> f (meet A B) <= meet (f A) (f B)"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    60
  by (auto simp add: mono_def)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    61
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    62
lemma Sup_insert[simp]: "Sup (insert (a::'a::comp_lat) A) = join a (Sup A)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    63
apply(simp add:Sup_def)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    64
apply(rule order_antisym)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    65
 apply(rule Meet_lower)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    66
 apply(clarsimp)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    67
 apply(rule le_joinI2)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    68
 apply(rule Meet_greatest)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    69
 apply blast
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    70
apply simp
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    71
apply rule
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    72
 apply(rule Meet_greatest)apply blast
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    73
apply(rule Meet_greatest)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    74
apply(rule Meet_lower)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    75
apply blast
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    76
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    77
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    78
lemma bot_least[simp]: "Sup{} \<le> (x::'a::comp_lat)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    79
apply(simp add: Sup_def)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    80
apply(rule Meet_lower)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    81
apply blast
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    82
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    83
(*
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    84
lemma Meet_singleton[simp]: "Meet{a} = (a::'a::comp_lat)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    85
apply(rule order_antisym)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    86
 apply(simp add: Meet_lower)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    87
apply(rule Meet_greatest)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    88
apply(simp)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    89
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    90
*)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    91
lemma le_SupI: "(l::'a::comp_lat) : M \<Longrightarrow> l \<le> Sup M"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    92
apply(simp add:Sup_def)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    93
apply(rule Meet_greatest)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    94
apply(simp)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    95
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    96
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    97
lemma le_SUPI: "(l::'a::comp_lat) = M i \<Longrightarrow> l \<le> (SUP i. M i)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    98
apply(simp add:SUP_def)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
    99
apply(blast intro:le_SupI)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   100
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   101
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   102
lemma Sup_leI: "(!!x. x:M \<Longrightarrow> x \<le> u) \<Longrightarrow> Sup M \<le> (u::'a::comp_lat)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   103
apply(simp add:Sup_def)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   104
apply(rule Meet_lower)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   105
apply(blast)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   106
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   107
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   108
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   109
lemma SUP_leI: "(!!i. M i \<le> u) \<Longrightarrow> (SUP i. M i) \<le> (u::'a::comp_lat)"
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   110
apply(simp add:SUP_def)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   111
apply(blast intro!:Sup_leI)
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   112
done
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   113
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   114
lemma SUP_const[simp]: "(SUP i. M) = (M::'a::comp_lat)"
21316
4d913b8bccf1 image_constant_conv no longer [simp]
nipkow
parents: 21312
diff changeset
   115
by(simp add:SUP_def image_constant_conv join_absorp1)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   116
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   117
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   118
subsection {* Some instances of the type class of complete lattices *}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   119
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   120
subsubsection {* Booleans *}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   121
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   122
instance bool :: ord ..
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   123
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   124
defs
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   125
  le_bool_def: "P <= Q == P \<longrightarrow> Q"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   126
  less_bool_def: "P < Q == (P::bool) <= Q \<and> P \<noteq> Q"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   127
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   128
theorem le_boolI: "(P \<Longrightarrow> Q) \<Longrightarrow> P <= Q"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   129
  by (simp add: le_bool_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   130
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   131
theorem le_boolI': "P \<longrightarrow> Q \<Longrightarrow> P <= Q"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   132
  by (simp add: le_bool_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   133
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   134
theorem le_boolE: "P <= Q \<Longrightarrow> P \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   135
  by (simp add: le_bool_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   136
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   137
theorem le_boolD: "P <= Q \<Longrightarrow> P \<longrightarrow> Q"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   138
  by (simp add: le_bool_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   139
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   140
instance bool :: order
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   141
  apply intro_classes
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   142
  apply (unfold le_bool_def less_bool_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   143
  apply iprover+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   144
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   145
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   146
defs Meet_bool_def: "Meet A == ALL x:A. x"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   147
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   148
instance bool :: comp_lat
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   149
  apply intro_classes
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   150
  apply (unfold Meet_bool_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   151
  apply (iprover intro!: le_boolI elim: ballE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   152
  apply (iprover intro!: ballI le_boolI elim: ballE le_boolE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   153
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   154
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   155
theorem meet_bool_eq: "meet P Q = (P \<and> Q)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   156
  apply (rule order_antisym)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   157
  apply (rule le_boolI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   158
  apply (rule conjI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   159
  apply (rule le_boolE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   160
  apply (rule meet_left_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   161
  apply assumption+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   162
  apply (rule le_boolE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   163
  apply (rule meet_right_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   164
  apply assumption+
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   165
  apply (rule le_meetI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   166
  apply (rule le_boolI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   167
  apply (erule conjunct1)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   168
  apply (rule le_boolI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   169
  apply (erule conjunct2)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   170
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   171
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   172
theorem join_bool_eq: "join P Q = (P \<or> Q)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   173
  apply (rule order_antisym)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   174
  apply (rule join_leI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   175
  apply (rule le_boolI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   176
  apply (erule disjI1)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   177
  apply (rule le_boolI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   178
  apply (erule disjI2)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   179
  apply (rule le_boolI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   180
  apply (erule disjE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   181
  apply (rule le_boolE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   182
  apply (rule join_left_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   183
  apply assumption+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   184
  apply (rule le_boolE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   185
  apply (rule join_right_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   186
  apply assumption+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   187
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   188
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   189
theorem Sup_bool_eq: "Sup A = (EX x:A. x)"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   190
  apply (rule order_antisym)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   191
  apply (rule Sup_least)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   192
  apply (rule le_boolI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   193
  apply (erule bexI, assumption)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   194
  apply (rule le_boolI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   195
  apply (erule bexE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   196
  apply (rule le_boolE)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   197
  apply (rule Sup_upper)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   198
  apply assumption+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   199
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   200
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   201
subsubsection {* Functions *}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   202
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   203
instance "fun" :: (type, ord) ord ..
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   204
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   205
defs
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   206
  le_fun_def: "f <= g == \<forall>x. f x <= g x"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   207
  less_fun_def: "f < g == (f::'a\<Rightarrow>'b::ord) <= g \<and> f \<noteq> g"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   208
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   209
theorem le_funI: "(\<And>x. f x <= g x) \<Longrightarrow> f <= g"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   210
  by (simp add: le_fun_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   211
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   212
theorem le_funE: "f <= g \<Longrightarrow> (f x <= g x \<Longrightarrow> P) \<Longrightarrow> P"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   213
  by (simp add: le_fun_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   214
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   215
theorem le_funD: "f <= g \<Longrightarrow> f x <= g x"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   216
  by (simp add: le_fun_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   217
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   218
text {*
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   219
Handy introduction and elimination rules for @{text "\<le>"}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   220
on unary and binary predicates
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   221
*}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   222
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   223
lemma predicate1I [intro]:
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   224
  assumes PQ: "\<And>x. P x \<Longrightarrow> Q x"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   225
  shows "P \<le> Q"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   226
  apply (rule le_funI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   227
  apply (rule le_boolI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   228
  apply (rule PQ)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   229
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   230
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   231
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   232
lemma predicate1D [elim]: "P \<le> Q \<Longrightarrow> P x \<Longrightarrow> Q x"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   233
  apply (erule le_funE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   234
  apply (erule le_boolE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   235
  apply assumption+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   236
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   237
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   238
lemma predicate2I [intro]:
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   239
  assumes PQ: "\<And>x y. P x y \<Longrightarrow> Q x y"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   240
  shows "P \<le> Q"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   241
  apply (rule le_funI)+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   242
  apply (rule le_boolI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   243
  apply (rule PQ)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   244
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   245
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   246
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   247
lemma predicate2D [elim]: "P \<le> Q \<Longrightarrow> P x y \<Longrightarrow> Q x y"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   248
  apply (erule le_funE)+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   249
  apply (erule le_boolE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   250
  apply assumption+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   251
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   252
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   253
instance "fun" :: (type, order) order
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   254
  apply intro_classes
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   255
  apply (rule le_funI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   256
  apply (rule order_refl)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   257
  apply (rule le_funI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   258
  apply (erule le_funE)+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   259
  apply (erule order_trans)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   260
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   261
  apply (rule ext)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   262
  apply (erule le_funE)+
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   263
  apply (erule order_antisym)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   264
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   265
  apply (simp add: less_fun_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   266
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   267
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   268
defs Meet_fun_def: "Meet A == (\<lambda>x. Meet {y. EX f:A. y = f x})"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   269
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   270
instance "fun" :: (type, comp_lat) comp_lat
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   271
  apply intro_classes
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   272
  apply (unfold Meet_fun_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   273
  apply (rule le_funI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   274
  apply (rule Meet_lower)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   275
  apply (rule CollectI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   276
  apply (rule bexI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   277
  apply (rule refl)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   278
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   279
  apply (rule le_funI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   280
  apply (rule Meet_greatest)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   281
  apply (erule CollectE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   282
  apply (erule bexE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   283
  apply (iprover elim: le_funE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   284
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   285
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   286
theorem meet_fun_eq: "meet f g = (\<lambda>x. meet (f x) (g x))"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   287
  apply (rule order_antisym)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   288
  apply (rule le_funI)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   289
  apply (rule le_meetI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   290
  apply (rule le_funD [OF meet_left_le])
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   291
  apply (rule le_funD [OF meet_right_le])
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   292
  apply (rule le_meetI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   293
  apply (rule le_funI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   294
  apply (rule meet_left_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   295
  apply (rule le_funI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   296
  apply (rule meet_right_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   297
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   298
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   299
theorem join_fun_eq: "join f g = (\<lambda>x. join (f x) (g x))"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   300
  apply (rule order_antisym)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   301
  apply (rule join_leI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   302
  apply (rule le_funI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   303
  apply (rule join_left_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   304
  apply (rule le_funI)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   305
  apply (rule join_right_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   306
  apply (rule le_funI)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   307
  apply (rule join_leI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   308
  apply (rule le_funD [OF join_left_le])
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   309
  apply (rule le_funD [OF join_right_le])
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   310
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   311
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   312
theorem Sup_fun_eq: "Sup A = (\<lambda>x. Sup {y::'a::comp_lat. EX f:A. y = f x})"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   313
  apply (rule order_antisym)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   314
  apply (rule Sup_least)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   315
  apply (rule le_funI)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   316
  apply (rule Sup_upper)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   317
  apply fast
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   318
  apply (rule le_funI)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   319
  apply (rule Sup_least)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   320
  apply (erule CollectE)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   321
  apply (erule bexE)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   322
  apply (drule le_funD [OF Sup_upper])
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   323
  apply simp
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   324
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   325
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   326
subsubsection {* Sets *}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   327
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   328
defs Meet_set_def: "Meet S == \<Inter>S"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   329
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   330
instance set :: (type) comp_lat
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   331
  by intro_classes (auto simp add: Meet_set_def)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   332
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   333
theorem meet_set_eq: "meet A B = A \<inter> B"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   334
  apply (rule subset_antisym)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   335
  apply (rule Int_greatest)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   336
  apply (rule meet_left_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   337
  apply (rule meet_right_le)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   338
  apply (rule le_meetI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   339
  apply (rule Int_lower1)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   340
  apply (rule Int_lower2)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   341
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   342
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   343
theorem join_set_eq: "join A B = A \<union> B"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   344
  apply (rule subset_antisym)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   345
  apply (rule join_leI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   346
  apply (rule Un_upper1)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   347
  apply (rule Un_upper2)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   348
  apply (rule Un_least)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   349
  apply (rule join_left_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   350
  apply (rule join_right_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   351
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   352
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   353
theorem Sup_set_eq: "Sup S = \<Union>S"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   354
  apply (rule subset_antisym)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   355
  apply (rule Sup_least)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   356
  apply (erule Union_upper)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   357
  apply (rule Union_least)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   358
  apply (erule Sup_upper)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   359
  done
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   360
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   361
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   362
subsection {* Least and greatest fixed points *}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   363
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   364
constdefs
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   365
  lfp :: "(('a::comp_lat) => 'a) => 'a"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   366
  "lfp f == Meet {u. f u <= u}"    --{*least fixed point*}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   367
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   368
  gfp :: "(('a::comp_lat) => 'a) => 'a"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   369
  "gfp f == Sup {u. u <= f u}"    --{*greatest fixed point*}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   370
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   371
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   372
subsection{*Proof of Knaster-Tarski Theorem using @{term lfp}*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   373
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   374
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   375
text{*@{term "lfp f"} is the least upper bound of 
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   376
      the set @{term "{u. f(u) \<le> u}"} *}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   377
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   378
lemma lfp_lowerbound: "f A \<le> A ==> lfp f \<le> A"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   379
  by (auto simp add: lfp_def intro: Meet_lower)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   380
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   381
lemma lfp_greatest: "(!!u. f u \<le> u ==> A \<le> u) ==> A \<le> lfp f"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   382
  by (auto simp add: lfp_def intro: Meet_greatest)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   383
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   384
lemma lfp_lemma2: "mono f ==> f (lfp f) \<le> lfp f"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   385
  by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   386
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   387
lemma lfp_lemma3: "mono f ==> lfp f \<le> f (lfp f)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   388
  by (iprover intro: lfp_lemma2 monoD lfp_lowerbound)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   389
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   390
lemma lfp_unfold: "mono f ==> lfp f = f (lfp f)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   391
  by (iprover intro: order_antisym lfp_lemma2 lfp_lemma3)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   392
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   393
subsection{*General induction rules for least fixed points*}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   394
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   395
theorem lfp_induct:
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   396
  assumes mono: "mono f" and ind: "f (meet (lfp f) P) <= P"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   397
  shows "lfp f <= P"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   398
proof -
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   399
  have "meet (lfp f) P <= lfp f" by (rule meet_left_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   400
  with mono have "f (meet (lfp f) P) <= f (lfp f)" ..
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   401
  also from mono have "f (lfp f) = lfp f" by (rule lfp_unfold [symmetric])
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   402
  finally have "f (meet (lfp f) P) <= lfp f" .
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   403
  from this and ind have "f (meet (lfp f) P) <= meet (lfp f) P" by (rule le_meetI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   404
  hence "lfp f <= meet (lfp f) P" by (rule lfp_lowerbound)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   405
  also have "meet (lfp f) P <= P" by (rule meet_right_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   406
  finally show ?thesis .
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   407
qed
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   408
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   409
lemma lfp_induct_set:
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   410
  assumes lfp: "a: lfp(f)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   411
      and mono: "mono(f)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   412
      and indhyp: "!!x. [| x: f(lfp(f) Int {x. P(x)}) |] ==> P(x)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   413
  shows "P(a)"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   414
  by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp])
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   415
    (auto simp: meet_set_eq intro: indhyp)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   416
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   417
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   418
text{*Version of induction for binary relations*}
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   419
lemmas lfp_induct2 =  lfp_induct_set [of "(a,b)", split_format (complete)]
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   420
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   421
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   422
lemma lfp_ordinal_induct: 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   423
  assumes mono: "mono f"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   424
  shows "[| !!S. P S ==> P(f S); !!M. !S:M. P S ==> P(Union M) |] 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   425
         ==> P(lfp f)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   426
apply(subgoal_tac "lfp f = Union{S. S \<subseteq> lfp f & P S}")
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   427
 apply (erule ssubst, simp) 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   428
apply(subgoal_tac "Union{S. S \<subseteq> lfp f & P S} \<subseteq> lfp f")
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   429
 prefer 2 apply blast
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   430
apply(rule equalityI)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   431
 prefer 2 apply assumption
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   432
apply(drule mono [THEN monoD])
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   433
apply (cut_tac mono [THEN lfp_unfold], simp)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   434
apply (rule lfp_lowerbound, auto) 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   435
done
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   436
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   437
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   438
text{*Definition forms of @{text lfp_unfold} and @{text lfp_induct}, 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   439
    to control unfolding*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   440
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   441
lemma def_lfp_unfold: "[| h==lfp(f);  mono(f) |] ==> h = f(h)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   442
by (auto intro!: lfp_unfold)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   443
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   444
lemma def_lfp_induct: 
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   445
    "[| A == lfp(f); mono(f);
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   446
        f (meet A P) \<le> P
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   447
     |] ==> A \<le> P"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   448
  by (blast intro: lfp_induct)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   449
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   450
lemma def_lfp_induct_set: 
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   451
    "[| A == lfp(f);  mono(f);   a:A;                    
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   452
        !!x. [| x: f(A Int {x. P(x)}) |] ==> P(x)         
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   453
     |] ==> P(a)"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   454
  by (blast intro: lfp_induct_set)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   455
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   456
(*Monotonicity of lfp!*)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   457
lemma lfp_mono: "(!!Z. f Z \<le> g Z) ==> lfp f \<le> lfp g"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   458
  by (rule lfp_lowerbound [THEN lfp_greatest], blast intro: order_trans)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   459
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   460
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   461
subsection{*Proof of Knaster-Tarski Theorem using @{term gfp}*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   462
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   463
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   464
text{*@{term "gfp f"} is the greatest lower bound of 
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   465
      the set @{term "{u. u \<le> f(u)}"} *}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   466
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   467
lemma gfp_upperbound: "X \<le> f X ==> X \<le> gfp f"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   468
  by (auto simp add: gfp_def intro: Sup_upper)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   469
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   470
lemma gfp_least: "(!!u. u \<le> f u ==> u \<le> X) ==> gfp f \<le> X"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   471
  by (auto simp add: gfp_def intro: Sup_least)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   472
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   473
lemma gfp_lemma2: "mono f ==> gfp f \<le> f (gfp f)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   474
  by (iprover intro: gfp_least order_trans monoD gfp_upperbound)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   475
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   476
lemma gfp_lemma3: "mono f ==> f (gfp f) \<le> gfp f"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   477
  by (iprover intro: gfp_lemma2 monoD gfp_upperbound)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   478
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   479
lemma gfp_unfold: "mono f ==> gfp f = f (gfp f)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   480
  by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   481
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   482
subsection{*Coinduction rules for greatest fixed points*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   483
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   484
text{*weak version*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   485
lemma weak_coinduct: "[| a: X;  X \<subseteq> f(X) |] ==> a : gfp(f)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   486
by (rule gfp_upperbound [THEN subsetD], auto)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   487
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   488
lemma weak_coinduct_image: "!!X. [| a : X; g`X \<subseteq> f (g`X) |] ==> g a : gfp f"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   489
apply (erule gfp_upperbound [THEN subsetD])
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   490
apply (erule imageI)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   491
done
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   492
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   493
lemma coinduct_lemma:
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   494
     "[| X \<le> f (join X (gfp f));  mono f |] ==> join X (gfp f) \<le> f (join X (gfp f))"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   495
  apply (frule gfp_lemma2)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   496
  apply (drule mono_join)
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   497
  apply (rule join_leI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   498
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   499
  apply (rule order_trans)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   500
  apply (rule order_trans)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   501
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   502
  apply (rule join_right_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   503
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   504
  done
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   505
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   506
text{*strong version, thanks to Coen and Frost*}
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   507
lemma coinduct_set: "[| mono(f);  a: X;  X \<subseteq> f(X Un gfp(f)) |] ==> a : gfp(f)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   508
by (blast intro: weak_coinduct [OF _ coinduct_lemma, simplified join_set_eq])
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   509
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   510
lemma coinduct: "[| mono(f); X \<le> f (join X (gfp f)) |] ==> X \<le> gfp(f)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   511
  apply (rule order_trans)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   512
  apply (rule join_left_le)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   513
  apply (erule gfp_upperbound [OF coinduct_lemma])
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   514
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   515
  done
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   516
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   517
lemma gfp_fun_UnI2: "[| mono(f);  a: gfp(f) |] ==> a: f(X Un gfp(f))"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   518
by (blast dest: gfp_lemma2 mono_Un)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   519
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   520
subsection{*Even Stronger Coinduction Rule, by Martin Coen*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   521
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   522
text{* Weakens the condition @{term "X \<subseteq> f(X)"} to one expressed using both
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   523
  @{term lfp} and @{term gfp}*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   524
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   525
lemma coinduct3_mono_lemma: "mono(f) ==> mono(%x. f(x) Un X Un B)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17006
diff changeset
   526
by (iprover intro: subset_refl monoI Un_mono monoD)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   527
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   528
lemma coinduct3_lemma:
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   529
     "[| X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)));  mono(f) |]
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   530
      ==> lfp(%x. f(x) Un X Un gfp(f)) \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)))"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   531
apply (rule subset_trans)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   532
apply (erule coinduct3_mono_lemma [THEN lfp_lemma3])
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   533
apply (rule Un_least [THEN Un_least])
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   534
apply (rule subset_refl, assumption)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   535
apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   536
apply (rule monoD, assumption)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   537
apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   538
done
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   539
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   540
lemma coinduct3: 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   541
  "[| mono(f);  a:X;  X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   542
apply (rule coinduct3_lemma [THEN [2] weak_coinduct])
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   543
apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst], auto)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   544
done
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   545
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   546
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   547
text{*Definition forms of @{text gfp_unfold} and @{text coinduct}, 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   548
    to control unfolding*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   549
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   550
lemma def_gfp_unfold: "[| A==gfp(f);  mono(f) |] ==> A = f(A)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   551
by (auto intro!: gfp_unfold)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   552
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   553
lemma def_coinduct:
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   554
     "[| A==gfp(f);  mono(f);  X \<le> f(join X A) |] ==> X \<le> A"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   555
by (iprover intro!: coinduct)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   556
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   557
lemma def_coinduct_set:
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   558
     "[| A==gfp(f);  mono(f);  a:X;  X \<subseteq> f(X Un A) |] ==> a: A"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   559
by (auto intro!: coinduct_set)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   560
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   561
(*The version used in the induction/coinduction package*)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   562
lemma def_Collect_coinduct:
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   563
    "[| A == gfp(%w. Collect(P(w)));  mono(%w. Collect(P(w)));   
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   564
        a: X;  !!z. z: X ==> P (X Un A) z |] ==>  
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   565
     a : A"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   566
apply (erule def_coinduct_set, auto) 
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   567
done
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   568
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   569
lemma def_coinduct3:
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   570
    "[| A==gfp(f); mono(f);  a:X;  X \<subseteq> f(lfp(%x. f(x) Un X Un A)) |] ==> a: A"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   571
by (auto intro!: coinduct3)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   572
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   573
text{*Monotonicity of @{term gfp}!*}
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   574
lemma gfp_mono: "(!!Z. f Z \<le> g Z) ==> gfp f \<le> gfp g"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   575
  by (rule gfp_upperbound [THEN gfp_least], blast intro: order_trans)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   576
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   577
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   578
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   579
ML
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   580
{*
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   581
val lfp_def = thm "lfp_def";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   582
val lfp_lowerbound = thm "lfp_lowerbound";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   583
val lfp_greatest = thm "lfp_greatest";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   584
val lfp_unfold = thm "lfp_unfold";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   585
val lfp_induct = thm "lfp_induct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   586
val lfp_induct2 = thm "lfp_induct2";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   587
val lfp_ordinal_induct = thm "lfp_ordinal_induct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   588
val def_lfp_unfold = thm "def_lfp_unfold";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   589
val def_lfp_induct = thm "def_lfp_induct";
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   590
val def_lfp_induct_set = thm "def_lfp_induct_set";
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   591
val lfp_mono = thm "lfp_mono";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   592
val gfp_def = thm "gfp_def";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   593
val gfp_upperbound = thm "gfp_upperbound";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   594
val gfp_least = thm "gfp_least";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   595
val gfp_unfold = thm "gfp_unfold";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   596
val weak_coinduct = thm "weak_coinduct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   597
val weak_coinduct_image = thm "weak_coinduct_image";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   598
val coinduct = thm "coinduct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   599
val gfp_fun_UnI2 = thm "gfp_fun_UnI2";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   600
val coinduct3 = thm "coinduct3";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   601
val def_gfp_unfold = thm "def_gfp_unfold";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   602
val def_coinduct = thm "def_coinduct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   603
val def_Collect_coinduct = thm "def_Collect_coinduct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   604
val def_coinduct3 = thm "def_coinduct3";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   605
val gfp_mono = thm "gfp_mono";
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   606
val le_funI = thm "le_funI";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   607
val le_boolI = thm "le_boolI";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   608
val le_boolI' = thm "le_boolI'";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   609
val meet_fun_eq = thm "meet_fun_eq";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   610
val meet_bool_eq = thm "meet_bool_eq";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   611
val le_funE = thm "le_funE";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   612
val le_boolE = thm "le_boolE";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   613
val le_boolD = thm "le_boolD";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   614
val le_bool_def = thm "le_bool_def";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   615
val le_fun_def = thm "le_fun_def";
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   616
*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   617
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   618
end