src/HOL/MicroJava/BV/Semilat.thy
author nipkow
Fri, 05 Jan 2001 18:48:18 +0100
changeset 10797 028d22926a41
parent 10496 f2d304bdf3cc
child 10918 9679326489cd
permissions -rw-r--r--
^^ -> ``` Univalent -> single_valued
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     1
(*  Title:      HOL/BCV/Semilat.thy
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     2
    ID:         $Id$
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     3
    Author:     Tobias Nipkow
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     4
    Copyright   2000 TUM
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     5
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     6
Semilattices
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     7
*)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     8
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     9
header "Semilattices"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    10
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    11
theory Semilat = Main:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    12
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    13
types 'a ord    = "'a => 'a => bool"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    14
      'a binop  = "'a => 'a => 'a"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    15
      'a sl     = "'a set * 'a ord * 'a binop"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    16
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    17
consts
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    18
 "@lesub"   :: "'a => 'a ord => 'a => bool" ("(_ /<='__ _)" [50, 1000, 51] 50)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    19
 "@lesssub" :: "'a => 'a ord => 'a => bool" ("(_ /<'__ _)" [50, 1000, 51] 50)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    20
defs
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    21
lesub_def:   "x <=_r y == r x y"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    22
lesssub_def: "x <_r y  == x <=_r y & x ~= y"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    23
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    24
consts
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    25
 "@plussub" :: "'a => ('a => 'b => 'c) => 'b => 'c" ("(_ /+'__ _)" [65, 1000, 66] 65)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    26
defs
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    27
plussub_def: "x +_f y == f x y"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    28
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    29
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    30
constdefs
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    31
 ord :: "('a*'a)set => 'a ord"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    32
"ord r == %x y. (x,y):r"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    33
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    34
 order :: "'a ord => bool"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    35
"order r == (!x. x <=_r x) &
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    36
            (!x y. x <=_r y & y <=_r x --> x=y) &
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    37
            (!x y z. x <=_r y & y <=_r z --> x <=_r z)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    38
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    39
 acc :: "'a ord => bool"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    40
"acc r == wf{(y,x) . x <_r y}"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    41
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    42
 top :: "'a ord => 'a => bool"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    43
"top r T == !x. x <=_r T"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    44
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    45
 closed :: "'a set => 'a binop => bool"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    46
"closed A f == !x:A. !y:A. x +_f y : A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    47
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    48
 semilat :: "'a sl => bool"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    49
"semilat == %(A,r,f). order r & closed A f &
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    50
                (!x:A. !y:A. x <=_r x +_f y)  &
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    51
                (!x:A. !y:A. y <=_r x +_f y)  &
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    52
                (!x:A. !y:A. !z:A. x <=_r z & y <=_r z --> x +_f y <=_r z)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    53
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    54
 is_ub :: "('a*'a)set => 'a => 'a => 'a => bool"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    55
"is_ub r x y u == (x,u):r & (y,u):r"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    56
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    57
 is_lub :: "('a*'a)set => 'a => 'a => 'a => bool"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    58
"is_lub r x y u == is_ub r x y u & (!z. is_ub r x y z --> (u,z):r)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    59
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    60
 some_lub :: "('a*'a)set => 'a => 'a => 'a"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    61
"some_lub r x y == SOME z. is_lub r x y z"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    62
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    63
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    64
lemma order_refl [simp, intro]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    65
  "order r ==> x <=_r x";
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    66
  by (simp add: order_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    67
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    68
lemma order_antisym:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    69
  "[| order r; x <=_r y; y <=_r x |] ==> x = y"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    70
apply (unfold order_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    71
apply (simp (no_asm_simp))  
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    72
done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    73
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    74
lemma order_trans:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    75
   "[| order r; x <=_r y; y <=_r z |] ==> x <=_r z"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    76
apply (unfold order_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    77
apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    78
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    79
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    80
lemma order_less_irrefl [intro, simp]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    81
   "order r ==> ~ x <_r x"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    82
apply (unfold order_def lesssub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    83
apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    84
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    85
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    86
lemma order_less_trans:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    87
  "[| order r; x <_r y; y <_r z |] ==> x <_r z"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    88
apply (unfold order_def lesssub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    89
apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    90
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    91
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    92
lemma topD [simp, intro]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    93
  "top r T ==> x <=_r T"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    94
  by (simp add: top_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    95
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    96
lemma top_le_conv [simp]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    97
  "[| order r; top r T |] ==> (T <=_r x) = (x = T)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    98
  by (blast intro: order_antisym)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    99
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   100
lemma semilat_Def:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   101
"semilat(A,r,f) == order r & closed A f & 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   102
                 (!x:A. !y:A. x <=_r x +_f y) & 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   103
                 (!x:A. !y:A. y <=_r x +_f y) & 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   104
                 (!x:A. !y:A. !z:A. x <=_r z & y <=_r z --> x +_f y <=_r z)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   105
apply (unfold semilat_def Product_Type.split [THEN eq_reflection])
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   106
apply (rule refl [THEN eq_reflection])
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   107
done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   108
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   109
lemma semilatDorderI [simp, intro]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   110
  "semilat(A,r,f) ==> order r"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   111
  by (simp add: semilat_Def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   112
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   113
lemma semilatDclosedI [simp, intro]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   114
  "semilat(A,r,f) ==> closed A f"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   115
apply (unfold semilat_Def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   116
apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   117
done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   118
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   119
lemma semilat_ub1 [simp]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   120
  "[| semilat(A,r,f); x:A; y:A |] ==> x <=_r x +_f y"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   121
  by (unfold semilat_Def, simp)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   122
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   123
lemma semilat_ub2 [simp]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   124
  "[| semilat(A,r,f); x:A; y:A |] ==> y <=_r x +_f y"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   125
  by (unfold semilat_Def, simp)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   126
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   127
lemma semilat_lub [simp]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   128
 "[| x <=_r z; y <=_r z; semilat(A,r,f); x:A; y:A; z:A |] ==> x +_f y <=_r z";
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   129
  by (unfold semilat_Def, simp)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   130
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   131
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   132
lemma plus_le_conv [simp]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   133
  "[| x:A; y:A; z:A; semilat(A,r,f) |] 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   134
  ==> (x +_f y <=_r z) = (x <=_r z & y <=_r z)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   135
apply (unfold semilat_Def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   136
apply (blast intro: semilat_ub1 semilat_ub2 semilat_lub order_trans)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   137
done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   138
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   139
lemma le_iff_plus_unchanged:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   140
  "[| x:A; y:A; semilat(A,r,f) |] ==> (x <=_r y) = (x +_f y = y)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   141
apply (rule iffI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   142
 apply (intro semilatDorderI order_antisym semilat_lub order_refl semilat_ub2, assumption+)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   143
apply (erule subst)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   144
apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   145
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   146
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   147
lemma le_iff_plus_unchanged2:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   148
  "[| x:A; y:A; semilat(A,r,f) |] ==> (x <=_r y) = (y +_f x = y)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   149
apply (rule iffI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   150
 apply (intro semilatDorderI order_antisym semilat_lub order_refl semilat_ub1, assumption+)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   151
apply (erule subst)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   152
apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   153
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   154
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   155
(*** closed ***)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   156
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   157
lemma closedD:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   158
  "[| closed A f; x:A; y:A |] ==> x +_f y : A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   159
apply (unfold closed_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   160
apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   161
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   162
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   163
lemma closed_UNIV [simp]: "closed UNIV f"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   164
  by (simp add: closed_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   165
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   166
(*** lub ***)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   167
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   168
lemma is_lubD:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   169
  "is_lub r x y u ==> is_ub r x y u & (!z. is_ub r x y z --> (u,z):r)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   170
  by (simp add: is_lub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   171
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   172
lemma is_ubI:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   173
  "[| (x,u) : r; (y,u) : r |] ==> is_ub r x y u"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   174
  by (simp add: is_ub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   175
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   176
lemma is_ubD:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   177
  "is_ub r x y u ==> (x,u) : r & (y,u) : r"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   178
  by (simp add: is_ub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   179
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   180
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   181
lemma is_lub_bigger1 [iff]:  
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   182
  "is_lub (r^* ) x y y = ((x,y):r^* )"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   183
apply (unfold is_lub_def is_ub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   184
apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   185
done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   186
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   187
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   188
lemma is_lub_bigger2 [iff]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   189
  "is_lub (r^* ) x y x = ((y,x):r^* )"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   190
apply (unfold is_lub_def is_ub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   191
apply blast 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   192
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   193
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   194
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   195
lemma extend_lub:
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10496
diff changeset
   196
  "[| single_valued r; is_lub (r^* ) x y u; (x',x) : r |] 
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   197
  ==> EX v. is_lub (r^* ) x' y v"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   198
apply (unfold is_lub_def is_ub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   199
apply (case_tac "(y,x) : r^*")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   200
 apply (case_tac "(y,x') : r^*")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   201
  apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   202
 apply (blast intro: r_into_rtrancl elim: converse_rtranclE
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10496
diff changeset
   203
               dest: single_valuedD)
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   204
apply (rule exI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   205
apply (rule conjI)
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10496
diff changeset
   206
 apply (blast intro: rtrancl_into_rtrancl2 dest: single_valuedD)
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   207
apply (blast intro: rtrancl_into_rtrancl rtrancl_into_rtrancl2 
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10496
diff changeset
   208
             elim: converse_rtranclE dest: single_valuedD)
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   209
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   210
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10496
diff changeset
   211
lemma single_valued_has_lubs [rule_format]:
028d22926a41 ^^ -> ```
nipkow
parents: 10496
diff changeset
   212
  "[| single_valued r; (x,u) : r^* |] ==> (!y. (y,u) : r^* --> 
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   213
  (EX z. is_lub (r^* ) x y z))"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   214
apply (erule converse_rtrancl_induct)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   215
 apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   216
 apply (erule converse_rtrancl_induct)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   217
  apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   218
 apply (blast intro: rtrancl_into_rtrancl2)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   219
apply (blast intro: extend_lub)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   220
done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   221
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   222
lemma some_lub_conv:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   223
  "[| acyclic r; is_lub (r^* ) x y u |] ==> some_lub (r^* ) x y = u"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   224
apply (unfold some_lub_def is_lub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   225
apply (rule someI2)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   226
 apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   227
apply (blast intro: antisymD dest!: acyclic_impl_antisym_rtrancl)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   228
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   229
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   230
lemma is_lub_some_lub:
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10496
diff changeset
   231
  "[| single_valued r; acyclic r; (x,u):r^*; (y,u):r^* |] 
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   232
  ==> is_lub (r^* ) x y (some_lub (r^* ) x y)";
10797
028d22926a41 ^^ -> ```
nipkow
parents: 10496
diff changeset
   233
  by (fastsimp dest: single_valued_has_lubs simp add: some_lub_conv)
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   234
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   235
end