src/HOL/Predicate.thy
author haftmann
Fri, 06 Mar 2009 20:30:17 +0100
changeset 30328 ab47f43f7581
parent 26797 a6cb51c314f2
child 30376 e8cc806a3755
permissions -rw-r--r--
added enumeration of predicates
berghofe@22259
     1
(*  Title:      HOL/Predicate.thy
haftmann@30328
     2
    Author:     Stefan Berghofer and Lukas Bulwahn and Florian Haftmann, TU Muenchen
berghofe@22259
     3
*)
berghofe@22259
     4
haftmann@30328
     5
header {* Predicates as relations and enumerations *}
berghofe@22259
     6
berghofe@22259
     7
theory Predicate
haftmann@23708
     8
imports Inductive Relation
berghofe@22259
     9
begin
berghofe@22259
    10
haftmann@30328
    11
notation
haftmann@30328
    12
  inf (infixl "\<sqinter>" 70) and
haftmann@30328
    13
  sup (infixl "\<squnion>" 65) and
haftmann@30328
    14
  Inf ("\<Sqinter>_" [900] 900) and
haftmann@30328
    15
  Sup ("\<Squnion>_" [900] 900) and
haftmann@30328
    16
  top ("\<top>") and
haftmann@30328
    17
  bot ("\<bottom>")
haftmann@30328
    18
haftmann@30328
    19
haftmann@30328
    20
subsection {* Predicates as (complete) lattices *}
haftmann@30328
    21
haftmann@30328
    22
subsubsection {* @{const sup} on @{typ bool} *}
haftmann@30328
    23
haftmann@30328
    24
lemma sup_boolI1:
haftmann@30328
    25
  "P \<Longrightarrow> P \<squnion> Q"
haftmann@30328
    26
  by (simp add: sup_bool_eq)
haftmann@30328
    27
haftmann@30328
    28
lemma sup_boolI2:
haftmann@30328
    29
  "Q \<Longrightarrow> P \<squnion> Q"
haftmann@30328
    30
  by (simp add: sup_bool_eq)
haftmann@30328
    31
haftmann@30328
    32
lemma sup_boolE:
haftmann@30328
    33
  "P \<squnion> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R"
haftmann@30328
    34
  by (auto simp add: sup_bool_eq)
haftmann@30328
    35
haftmann@30328
    36
haftmann@30328
    37
subsubsection {* Equality and Subsets *}
berghofe@22259
    38
berghofe@26797
    39
lemma pred_equals_eq: "((\<lambda>x. x \<in> R) = (\<lambda>x. x \<in> S)) = (R = S)"
berghofe@26797
    40
  by (simp add: mem_def)
berghofe@22259
    41
berghofe@23741
    42
lemma pred_equals_eq2 [pred_set_conv]: "((\<lambda>x y. (x, y) \<in> R) = (\<lambda>x y. (x, y) \<in> S)) = (R = S)"
berghofe@26797
    43
  by (simp add: expand_fun_eq mem_def)
berghofe@22259
    44
berghofe@26797
    45
lemma pred_subset_eq: "((\<lambda>x. x \<in> R) <= (\<lambda>x. x \<in> S)) = (R <= S)"
berghofe@26797
    46
  by (simp add: mem_def)
berghofe@22259
    47
berghofe@23741
    48
lemma pred_subset_eq2 [pred_set_conv]: "((\<lambda>x y. (x, y) \<in> R) <= (\<lambda>x y. (x, y) \<in> S)) = (R <= S)"
berghofe@22259
    49
  by fast
berghofe@22259
    50
berghofe@22259
    51
haftmann@30328
    52
subsubsection {* Top and bottom elements *}
berghofe@22259
    53
berghofe@23741
    54
lemma top1I [intro!]: "top x"
berghofe@23741
    55
  by (simp add: top_fun_eq top_bool_eq)
berghofe@22259
    56
berghofe@23741
    57
lemma top2I [intro!]: "top x y"
berghofe@23741
    58
  by (simp add: top_fun_eq top_bool_eq)
berghofe@22259
    59
berghofe@23741
    60
lemma bot1E [elim!]: "bot x \<Longrightarrow> P"
berghofe@23741
    61
  by (simp add: bot_fun_eq bot_bool_eq)
berghofe@22259
    62
berghofe@23741
    63
lemma bot2E [elim!]: "bot x y \<Longrightarrow> P"
berghofe@23741
    64
  by (simp add: bot_fun_eq bot_bool_eq)
berghofe@23741
    65
berghofe@23741
    66
haftmann@30328
    67
subsubsection {* The empty set *}
berghofe@23741
    68
berghofe@23741
    69
lemma bot_empty_eq: "bot = (\<lambda>x. x \<in> {})"
berghofe@23741
    70
  by (auto simp add: expand_fun_eq)
berghofe@23741
    71
berghofe@23741
    72
lemma bot_empty_eq2: "bot = (\<lambda>x y. (x, y) \<in> {})"
berghofe@23741
    73
  by (auto simp add: expand_fun_eq)
berghofe@23741
    74
berghofe@23741
    75
haftmann@30328
    76
subsubsection {* Binary union *}
berghofe@22259
    77
haftmann@22422
    78
lemma sup1_iff [simp]: "sup A B x \<longleftrightarrow> A x | B x"
haftmann@22422
    79
  by (simp add: sup_fun_eq sup_bool_eq)
berghofe@22259
    80
haftmann@22422
    81
lemma sup2_iff [simp]: "sup A B x y \<longleftrightarrow> A x y | B x y"
haftmann@22422
    82
  by (simp add: sup_fun_eq sup_bool_eq)
berghofe@22259
    83
berghofe@23741
    84
lemma sup_Un_eq [pred_set_conv]: "sup (\<lambda>x. x \<in> R) (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<union> S)"
berghofe@23741
    85
  by (simp add: expand_fun_eq)
berghofe@23741
    86
berghofe@23741
    87
lemma sup_Un_eq2 [pred_set_conv]: "sup (\<lambda>x y. (x, y) \<in> R) (\<lambda>x y. (x, y) \<in> S) = (\<lambda>x y. (x, y) \<in> R \<union> S)"
berghofe@23741
    88
  by (simp add: expand_fun_eq)
berghofe@23741
    89
haftmann@22422
    90
lemma sup1I1 [elim?]: "A x \<Longrightarrow> sup A B x"
berghofe@22259
    91
  by simp
berghofe@22259
    92
haftmann@22422
    93
lemma sup2I1 [elim?]: "A x y \<Longrightarrow> sup A B x y"
berghofe@22259
    94
  by simp
berghofe@22259
    95
berghofe@23741
    96
lemma sup1I2 [elim?]: "B x \<Longrightarrow> sup A B x"
berghofe@22259
    97
  by simp
berghofe@22259
    98
berghofe@23741
    99
lemma sup2I2 [elim?]: "B x y \<Longrightarrow> sup A B x y"
berghofe@22259
   100
  by simp
berghofe@22259
   101
berghofe@22259
   102
text {*
berghofe@22259
   103
  \medskip Classical introduction rule: no commitment to @{text A} vs
berghofe@22259
   104
  @{text B}.
berghofe@22259
   105
*}
berghofe@22259
   106
haftmann@22422
   107
lemma sup1CI [intro!]: "(~ B x ==> A x) ==> sup A B x"
berghofe@22259
   108
  by auto
berghofe@22259
   109
haftmann@22422
   110
lemma sup2CI [intro!]: "(~ B x y ==> A x y) ==> sup A B x y"
berghofe@22259
   111
  by auto
berghofe@22259
   112
haftmann@22422
   113
lemma sup1E [elim!]: "sup A B x ==> (A x ==> P) ==> (B x ==> P) ==> P"
berghofe@22259
   114
  by simp iprover
berghofe@22259
   115
haftmann@22422
   116
lemma sup2E [elim!]: "sup A B x y ==> (A x y ==> P) ==> (B x y ==> P) ==> P"
berghofe@22259
   117
  by simp iprover
berghofe@22259
   118
berghofe@22259
   119
haftmann@30328
   120
subsubsection {* Binary intersection *}
berghofe@22259
   121
haftmann@22422
   122
lemma inf1_iff [simp]: "inf A B x \<longleftrightarrow> A x \<and> B x"
haftmann@22422
   123
  by (simp add: inf_fun_eq inf_bool_eq)
berghofe@22259
   124
haftmann@22422
   125
lemma inf2_iff [simp]: "inf A B x y \<longleftrightarrow> A x y \<and> B x y"
haftmann@22422
   126
  by (simp add: inf_fun_eq inf_bool_eq)
berghofe@22259
   127
berghofe@23741
   128
lemma inf_Int_eq [pred_set_conv]: "inf (\<lambda>x. x \<in> R) (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<inter> S)"
berghofe@23741
   129
  by (simp add: expand_fun_eq)
berghofe@23741
   130
berghofe@23741
   131
lemma inf_Int_eq2 [pred_set_conv]: "inf (\<lambda>x y. (x, y) \<in> R) (\<lambda>x y. (x, y) \<in> S) = (\<lambda>x y. (x, y) \<in> R \<inter> S)"
berghofe@23741
   132
  by (simp add: expand_fun_eq)
berghofe@23741
   133
haftmann@22422
   134
lemma inf1I [intro!]: "A x ==> B x ==> inf A B x"
berghofe@22259
   135
  by simp
berghofe@22259
   136
haftmann@22422
   137
lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y"
berghofe@22259
   138
  by simp
berghofe@22259
   139
haftmann@22422
   140
lemma inf1D1: "inf A B x ==> A x"
berghofe@22259
   141
  by simp
berghofe@22259
   142
haftmann@22422
   143
lemma inf2D1: "inf A B x y ==> A x y"
berghofe@22259
   144
  by simp
berghofe@22259
   145
haftmann@22422
   146
lemma inf1D2: "inf A B x ==> B x"
berghofe@22259
   147
  by simp
berghofe@22259
   148
haftmann@22422
   149
lemma inf2D2: "inf A B x y ==> B x y"
berghofe@22259
   150
  by simp
berghofe@22259
   151
haftmann@22422
   152
lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P"
berghofe@22259
   153
  by simp
berghofe@22259
   154
haftmann@22422
   155
lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P"
berghofe@22259
   156
  by simp
berghofe@22259
   157
berghofe@22259
   158
haftmann@30328
   159
subsubsection {* Unions of families *}
berghofe@22259
   160
berghofe@22430
   161
lemma SUP1_iff [simp]: "(SUP x:A. B x) b = (EX x:A. B x b)"
haftmann@24345
   162
  by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast
berghofe@22259
   163
berghofe@22430
   164
lemma SUP2_iff [simp]: "(SUP x:A. B x) b c = (EX x:A. B x b c)"
haftmann@24345
   165
  by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast
berghofe@22259
   166
berghofe@22430
   167
lemma SUP1_I [intro]: "a : A ==> B a b ==> (SUP x:A. B x) b"
berghofe@22259
   168
  by auto
berghofe@22259
   169
berghofe@22430
   170
lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c"
berghofe@22259
   171
  by auto
berghofe@22259
   172
berghofe@22430
   173
lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R"
berghofe@22430
   174
  by auto
berghofe@22259
   175
berghofe@22430
   176
lemma SUP2_E [elim!]: "(SUP x:A. B x) b c ==> (!!x. x : A ==> B x b c ==> R) ==> R"
berghofe@22430
   177
  by auto
berghofe@22430
   178
berghofe@23741
   179
lemma SUP_UN_eq: "(SUP i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (UN i. r i))"
berghofe@23741
   180
  by (simp add: expand_fun_eq)
berghofe@22430
   181
berghofe@23741
   182
lemma SUP_UN_eq2: "(SUP i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (UN i. r i))"
berghofe@23741
   183
  by (simp add: expand_fun_eq)
berghofe@22430
   184
berghofe@22430
   185
haftmann@30328
   186
subsubsection {* Intersections of families *}
berghofe@22430
   187
berghofe@22430
   188
lemma INF1_iff [simp]: "(INF x:A. B x) b = (ALL x:A. B x b)"
berghofe@22430
   189
  by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast
berghofe@22430
   190
berghofe@22430
   191
lemma INF2_iff [simp]: "(INF x:A. B x) b c = (ALL x:A. B x b c)"
berghofe@22430
   192
  by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast
berghofe@22430
   193
berghofe@22430
   194
lemma INF1_I [intro!]: "(!!x. x : A ==> B x b) ==> (INF x:A. B x) b"
berghofe@22430
   195
  by auto
berghofe@22430
   196
berghofe@22430
   197
lemma INF2_I [intro!]: "(!!x. x : A ==> B x b c) ==> (INF x:A. B x) b c"
berghofe@22430
   198
  by auto
berghofe@22430
   199
berghofe@22430
   200
lemma INF1_D [elim]: "(INF x:A. B x) b ==> a : A ==> B a b"
berghofe@22430
   201
  by auto
berghofe@22430
   202
berghofe@22430
   203
lemma INF2_D [elim]: "(INF x:A. B x) b c ==> a : A ==> B a b c"
berghofe@22430
   204
  by auto
berghofe@22430
   205
berghofe@22430
   206
lemma INF1_E [elim]: "(INF x:A. B x) b ==> (B a b ==> R) ==> (a ~: A ==> R) ==> R"
berghofe@22430
   207
  by auto
berghofe@22430
   208
berghofe@22430
   209
lemma INF2_E [elim]: "(INF x:A. B x) b c ==> (B a b c ==> R) ==> (a ~: A ==> R) ==> R"
berghofe@22430
   210
  by auto
berghofe@22259
   211
berghofe@23741
   212
lemma INF_INT_eq: "(INF i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (INT i. r i))"
berghofe@23741
   213
  by (simp add: expand_fun_eq)
berghofe@23741
   214
berghofe@23741
   215
lemma INF_INT_eq2: "(INF i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (INT i. r i))"
berghofe@23741
   216
  by (simp add: expand_fun_eq)
berghofe@23741
   217
berghofe@22259
   218
haftmann@30328
   219
subsection {* Predicates as relations *}
haftmann@30328
   220
haftmann@30328
   221
subsubsection {* Composition  *}
berghofe@22259
   222
berghofe@23741
   223
inductive
berghofe@22259
   224
  pred_comp  :: "['b => 'c => bool, 'a => 'b => bool] => 'a => 'c => bool"
berghofe@22259
   225
    (infixr "OO" 75)
berghofe@22259
   226
  for r :: "'b => 'c => bool" and s :: "'a => 'b => bool"
berghofe@22259
   227
where
berghofe@22259
   228
  pred_compI [intro]: "s a b ==> r b c ==> (r OO s) a c"
berghofe@22259
   229
berghofe@23741
   230
inductive_cases pred_compE [elim!]: "(r OO s) a c"
berghofe@22259
   231
berghofe@22259
   232
lemma pred_comp_rel_comp_eq [pred_set_conv]:
berghofe@23741
   233
  "((\<lambda>x y. (x, y) \<in> r) OO (\<lambda>x y. (x, y) \<in> s)) = (\<lambda>x y. (x, y) \<in> r O s)"
berghofe@22259
   234
  by (auto simp add: expand_fun_eq elim: pred_compE)
berghofe@22259
   235
berghofe@22259
   236
haftmann@30328
   237
subsubsection {* Converse *}
berghofe@22259
   238
berghofe@23741
   239
inductive
berghofe@22259
   240
  conversep :: "('a => 'b => bool) => 'b => 'a => bool"
berghofe@22259
   241
    ("(_^--1)" [1000] 1000)
berghofe@22259
   242
  for r :: "'a => 'b => bool"
berghofe@22259
   243
where
berghofe@22259
   244
  conversepI: "r a b ==> r^--1 b a"
berghofe@22259
   245
berghofe@22259
   246
notation (xsymbols)
berghofe@22259
   247
  conversep  ("(_\<inverse>\<inverse>)" [1000] 1000)
berghofe@22259
   248
berghofe@22259
   249
lemma conversepD:
berghofe@22259
   250
  assumes ab: "r^--1 a b"
berghofe@22259
   251
  shows "r b a" using ab
berghofe@22259
   252
  by cases simp
berghofe@22259
   253
berghofe@22259
   254
lemma conversep_iff [iff]: "r^--1 a b = r b a"
berghofe@22259
   255
  by (iprover intro: conversepI dest: conversepD)
berghofe@22259
   256
berghofe@22259
   257
lemma conversep_converse_eq [pred_set_conv]:
berghofe@23741
   258
  "(\<lambda>x y. (x, y) \<in> r)^--1 = (\<lambda>x y. (x, y) \<in> r^-1)"
berghofe@22259
   259
  by (auto simp add: expand_fun_eq)
berghofe@22259
   260
berghofe@22259
   261
lemma conversep_conversep [simp]: "(r^--1)^--1 = r"
berghofe@22259
   262
  by (iprover intro: order_antisym conversepI dest: conversepD)
berghofe@22259
   263
berghofe@22259
   264
lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1"
berghofe@22259
   265
  by (iprover intro: order_antisym conversepI pred_compI
berghofe@22259
   266
    elim: pred_compE dest: conversepD)
berghofe@22259
   267
haftmann@22422
   268
lemma converse_meet: "(inf r s)^--1 = inf r^--1 s^--1"
haftmann@22422
   269
  by (simp add: inf_fun_eq inf_bool_eq)
berghofe@22259
   270
    (iprover intro: conversepI ext dest: conversepD)
berghofe@22259
   271
haftmann@22422
   272
lemma converse_join: "(sup r s)^--1 = sup r^--1 s^--1"
haftmann@22422
   273
  by (simp add: sup_fun_eq sup_bool_eq)
berghofe@22259
   274
    (iprover intro: conversepI ext dest: conversepD)
berghofe@22259
   275
berghofe@22259
   276
lemma conversep_noteq [simp]: "(op ~=)^--1 = op ~="
berghofe@22259
   277
  by (auto simp add: expand_fun_eq)
berghofe@22259
   278
berghofe@22259
   279
lemma conversep_eq [simp]: "(op =)^--1 = op ="
berghofe@22259
   280
  by (auto simp add: expand_fun_eq)
berghofe@22259
   281
berghofe@22259
   282
haftmann@30328
   283
subsubsection {* Domain *}
berghofe@22259
   284
berghofe@23741
   285
inductive
berghofe@22259
   286
  DomainP :: "('a => 'b => bool) => 'a => bool"
berghofe@22259
   287
  for r :: "'a => 'b => bool"
berghofe@22259
   288
where
berghofe@22259
   289
  DomainPI [intro]: "r a b ==> DomainP r a"
berghofe@22259
   290
berghofe@23741
   291
inductive_cases DomainPE [elim!]: "DomainP r a"
berghofe@22259
   292
berghofe@23741
   293
lemma DomainP_Domain_eq [pred_set_conv]: "DomainP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Domain r)"
berghofe@26797
   294
  by (blast intro!: Orderings.order_antisym predicate1I)
berghofe@22259
   295
berghofe@22259
   296
haftmann@30328
   297
subsubsection {* Range *}
berghofe@22259
   298
berghofe@23741
   299
inductive
berghofe@22259
   300
  RangeP :: "('a => 'b => bool) => 'b => bool"
berghofe@22259
   301
  for r :: "'a => 'b => bool"
berghofe@22259
   302
where
berghofe@22259
   303
  RangePI [intro]: "r a b ==> RangeP r b"
berghofe@22259
   304
berghofe@23741
   305
inductive_cases RangePE [elim!]: "RangeP r b"
berghofe@22259
   306
berghofe@23741
   307
lemma RangeP_Range_eq [pred_set_conv]: "RangeP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Range r)"
berghofe@26797
   308
  by (blast intro!: Orderings.order_antisym predicate1I)
berghofe@22259
   309
berghofe@22259
   310
haftmann@30328
   311
subsubsection {* Inverse image *}
berghofe@22259
   312
berghofe@22259
   313
definition
berghofe@22259
   314
  inv_imagep :: "('b => 'b => bool) => ('a => 'b) => 'a => 'a => bool" where
berghofe@22259
   315
  "inv_imagep r f == %x y. r (f x) (f y)"
berghofe@22259
   316
berghofe@23741
   317
lemma [pred_set_conv]: "inv_imagep (\<lambda>x y. (x, y) \<in> r) f = (\<lambda>x y. (x, y) \<in> inv_image r f)"
berghofe@22259
   318
  by (simp add: inv_image_def inv_imagep_def)
berghofe@22259
   319
berghofe@22259
   320
lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)"
berghofe@22259
   321
  by (simp add: inv_imagep_def)
berghofe@22259
   322
berghofe@22259
   323
haftmann@30328
   324
subsubsection {* Powerset *}
berghofe@23741
   325
berghofe@23741
   326
definition Powp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool" where
berghofe@23741
   327
  "Powp A == \<lambda>B. \<forall>x \<in> B. A x"
berghofe@23741
   328
berghofe@23741
   329
lemma Powp_Pow_eq [pred_set_conv]: "Powp (\<lambda>x. x \<in> A) = (\<lambda>x. x \<in> Pow A)"
berghofe@23741
   330
  by (auto simp add: Powp_def expand_fun_eq)
berghofe@23741
   331
berghofe@26797
   332
lemmas Powp_mono [mono] = Pow_mono [to_pred pred_subset_eq]
berghofe@26797
   333
berghofe@23741
   334
haftmann@30328
   335
subsubsection {* Properties of relations *}
berghofe@22259
   336
berghofe@22259
   337
abbreviation antisymP :: "('a => 'a => bool) => bool" where
berghofe@23741
   338
  "antisymP r == antisym {(x, y). r x y}"
berghofe@22259
   339
berghofe@22259
   340
abbreviation transP :: "('a => 'a => bool) => bool" where
berghofe@23741
   341
  "transP r == trans {(x, y). r x y}"
berghofe@22259
   342
berghofe@22259
   343
abbreviation single_valuedP :: "('a => 'b => bool) => bool" where
berghofe@23741
   344
  "single_valuedP r == single_valued {(x, y). r x y}"
berghofe@22259
   345
haftmann@30328
   346
haftmann@30328
   347
subsection {* Predicates as enumerations *}
haftmann@30328
   348
haftmann@30328
   349
subsubsection {* The type of predicate enumerations (a monad) *}
haftmann@30328
   350
haftmann@30328
   351
datatype 'a pred = Pred "'a \<Rightarrow> bool"
haftmann@30328
   352
haftmann@30328
   353
primrec eval :: "'a pred \<Rightarrow> 'a \<Rightarrow> bool" where
haftmann@30328
   354
  eval_pred: "eval (Pred f) = f"
haftmann@30328
   355
haftmann@30328
   356
lemma Pred_eval [simp]:
haftmann@30328
   357
  "Pred (eval x) = x"
haftmann@30328
   358
  by (cases x) simp
haftmann@30328
   359
haftmann@30328
   360
lemma eval_inject: "eval x = eval y \<longleftrightarrow> x = y"
haftmann@30328
   361
  by (cases x) auto
haftmann@30328
   362
haftmann@30328
   363
definition single :: "'a \<Rightarrow> 'a pred" where
haftmann@30328
   364
  "single x = Pred ((op =) x)"
haftmann@30328
   365
haftmann@30328
   366
definition bind :: "'a pred \<Rightarrow> ('a \<Rightarrow> 'b pred) \<Rightarrow> 'b pred" (infixl "\<guillemotright>=" 70) where
haftmann@30328
   367
  "P \<guillemotright>= f = Pred (\<lambda>x. (\<exists>y. eval P y \<and> eval (f y) x))"
haftmann@30328
   368
haftmann@30328
   369
instantiation pred :: (type) complete_lattice
haftmann@30328
   370
begin
haftmann@30328
   371
haftmann@30328
   372
definition
haftmann@30328
   373
  "P \<le> Q \<longleftrightarrow> eval P \<le> eval Q"
haftmann@30328
   374
haftmann@30328
   375
definition
haftmann@30328
   376
  "P < Q \<longleftrightarrow> eval P < eval Q"
haftmann@30328
   377
haftmann@30328
   378
definition
haftmann@30328
   379
  "\<bottom> = Pred \<bottom>"
haftmann@30328
   380
haftmann@30328
   381
definition
haftmann@30328
   382
  "\<top> = Pred \<top>"
haftmann@30328
   383
haftmann@30328
   384
definition
haftmann@30328
   385
  "P \<sqinter> Q = Pred (eval P \<sqinter> eval Q)"
haftmann@30328
   386
haftmann@30328
   387
definition
haftmann@30328
   388
  "P \<squnion> Q = Pred (eval P \<squnion> eval Q)"
haftmann@30328
   389
haftmann@30328
   390
definition
haftmann@30328
   391
  "\<Sqinter>A = Pred (INFI A eval)"
haftmann@30328
   392
haftmann@30328
   393
definition
haftmann@30328
   394
  "\<Squnion>A = Pred (SUPR A eval)"
haftmann@30328
   395
haftmann@30328
   396
instance by default
haftmann@30328
   397
  (auto simp add: less_eq_pred_def less_pred_def
haftmann@30328
   398
    inf_pred_def sup_pred_def bot_pred_def top_pred_def
haftmann@30328
   399
    Inf_pred_def Sup_pred_def,
haftmann@30328
   400
    auto simp add: le_fun_def less_fun_def le_bool_def less_bool_def
haftmann@30328
   401
    eval_inject mem_def)
haftmann@30328
   402
berghofe@22259
   403
end
haftmann@30328
   404
haftmann@30328
   405
lemma bind_bind:
haftmann@30328
   406
  "(P \<guillemotright>= Q) \<guillemotright>= R = P \<guillemotright>= (\<lambda>x. Q x \<guillemotright>= R)"
haftmann@30328
   407
  by (auto simp add: bind_def expand_fun_eq)
haftmann@30328
   408
haftmann@30328
   409
lemma bind_single:
haftmann@30328
   410
  "P \<guillemotright>= single = P"
haftmann@30328
   411
  by (simp add: bind_def single_def)
haftmann@30328
   412
haftmann@30328
   413
lemma single_bind:
haftmann@30328
   414
  "single x \<guillemotright>= P = P x"
haftmann@30328
   415
  by (simp add: bind_def single_def)
haftmann@30328
   416
haftmann@30328
   417
lemma bottom_bind:
haftmann@30328
   418
  "\<bottom> \<guillemotright>= P = \<bottom>"
haftmann@30328
   419
  by (auto simp add: bot_pred_def bind_def expand_fun_eq)
haftmann@30328
   420
haftmann@30328
   421
lemma sup_bind:
haftmann@30328
   422
  "(P \<squnion> Q) \<guillemotright>= R = P \<guillemotright>= R \<squnion> Q \<guillemotright>= R"
haftmann@30328
   423
  by (auto simp add: bind_def sup_pred_def expand_fun_eq)
haftmann@30328
   424
haftmann@30328
   425
lemma Sup_bind: "(\<Squnion>A \<guillemotright>= f) = \<Squnion>((\<lambda>x. x \<guillemotright>= f) ` A)"
haftmann@30328
   426
  by (auto simp add: bind_def Sup_pred_def expand_fun_eq)
haftmann@30328
   427
haftmann@30328
   428
lemma pred_iffI:
haftmann@30328
   429
  assumes "\<And>x. eval A x \<Longrightarrow> eval B x"
haftmann@30328
   430
  and "\<And>x. eval B x \<Longrightarrow> eval A x"
haftmann@30328
   431
  shows "A = B"
haftmann@30328
   432
proof -
haftmann@30328
   433
  from assms have "\<And>x. eval A x \<longleftrightarrow> eval B x" by blast
haftmann@30328
   434
  then show ?thesis by (cases A, cases B) (simp add: expand_fun_eq)
haftmann@30328
   435
qed
haftmann@30328
   436
  
haftmann@30328
   437
lemma singleI: "eval (single x) x"
haftmann@30328
   438
  unfolding single_def by simp
haftmann@30328
   439
haftmann@30328
   440
lemma singleI_unit: "eval (single ()) x"
haftmann@30328
   441
  by simp (rule singleI)
haftmann@30328
   442
haftmann@30328
   443
lemma singleE: "eval (single x) y \<Longrightarrow> (y = x \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   444
  unfolding single_def by simp
haftmann@30328
   445
haftmann@30328
   446
lemma singleE': "eval (single x) y \<Longrightarrow> (x = y \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   447
  by (erule singleE) simp
haftmann@30328
   448
haftmann@30328
   449
lemma bindI: "eval P x \<Longrightarrow> eval (Q x) y \<Longrightarrow> eval (P \<guillemotright>= Q) y"
haftmann@30328
   450
  unfolding bind_def by auto
haftmann@30328
   451
haftmann@30328
   452
lemma bindE: "eval (R \<guillemotright>= Q) y \<Longrightarrow> (\<And>x. eval R x \<Longrightarrow> eval (Q x) y \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   453
  unfolding bind_def by auto
haftmann@30328
   454
haftmann@30328
   455
lemma botE: "eval \<bottom> x \<Longrightarrow> P"
haftmann@30328
   456
  unfolding bot_pred_def by auto
haftmann@30328
   457
haftmann@30328
   458
lemma supI1: "eval A x \<Longrightarrow> eval (A \<squnion> B) x"
haftmann@30328
   459
  unfolding sup_pred_def by simp
haftmann@30328
   460
haftmann@30328
   461
lemma supI2: "eval B x \<Longrightarrow> eval (A \<squnion> B) x" 
haftmann@30328
   462
  unfolding sup_pred_def by simp
haftmann@30328
   463
haftmann@30328
   464
lemma supE: "eval (A \<squnion> B) x \<Longrightarrow> (eval A x \<Longrightarrow> P) \<Longrightarrow> (eval B x \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   465
  unfolding sup_pred_def by auto
haftmann@30328
   466
haftmann@30328
   467
haftmann@30328
   468
subsubsection {* Derived operations *}
haftmann@30328
   469
haftmann@30328
   470
definition if_pred :: "bool \<Rightarrow> unit pred" where
haftmann@30328
   471
  if_pred_eq: "if_pred b = (if b then single () else \<bottom>)"
haftmann@30328
   472
haftmann@30328
   473
definition eq_pred :: "'a \<Rightarrow> 'a \<Rightarrow> unit pred" where
haftmann@30328
   474
  eq_pred_eq: "eq_pred a b = if_pred (a = b)"
haftmann@30328
   475
haftmann@30328
   476
definition not_pred :: "unit pred \<Rightarrow> unit pred" where
haftmann@30328
   477
  not_pred_eq: "not_pred P = (if eval P () then \<bottom> else single ())"
haftmann@30328
   478
haftmann@30328
   479
lemma if_predI: "P \<Longrightarrow> eval (if_pred P) ()"
haftmann@30328
   480
  unfolding if_pred_eq by (auto intro: singleI)
haftmann@30328
   481
haftmann@30328
   482
lemma if_predE: "eval (if_pred b) x \<Longrightarrow> (b \<Longrightarrow> x = () \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   483
  unfolding if_pred_eq by (cases b) (auto elim: botE)
haftmann@30328
   484
haftmann@30328
   485
lemma eq_predI: "eval (eq_pred a a) ()"
haftmann@30328
   486
  unfolding eq_pred_eq if_pred_eq by (auto intro: singleI)
haftmann@30328
   487
haftmann@30328
   488
lemma eq_predE: "eval (eq_pred a b) x \<Longrightarrow> (a = b \<Longrightarrow> x = () \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   489
  unfolding eq_pred_eq by (erule if_predE)
haftmann@30328
   490
haftmann@30328
   491
lemma not_predI: "\<not> P \<Longrightarrow> eval (not_pred (Pred (\<lambda>u. P))) ()"
haftmann@30328
   492
  unfolding not_pred_eq eval_pred by (auto intro: singleI)
haftmann@30328
   493
haftmann@30328
   494
lemma not_predI': "\<not> eval P () \<Longrightarrow> eval (not_pred P) ()"
haftmann@30328
   495
  unfolding not_pred_eq by (auto intro: singleI)
haftmann@30328
   496
haftmann@30328
   497
lemma not_predE: "eval (not_pred (Pred (\<lambda>u. P))) x \<Longrightarrow> (\<not> P \<Longrightarrow> thesis) \<Longrightarrow> thesis"
haftmann@30328
   498
  unfolding not_pred_eq
haftmann@30328
   499
  by (auto split: split_if_asm elim: botE)
haftmann@30328
   500
haftmann@30328
   501
lemma not_predE': "eval (not_pred P) x \<Longrightarrow> (\<not> eval P x \<Longrightarrow> thesis) \<Longrightarrow> thesis"
haftmann@30328
   502
  unfolding not_pred_eq
haftmann@30328
   503
  by (auto split: split_if_asm elim: botE)
haftmann@30328
   504
haftmann@30328
   505
haftmann@30328
   506
subsubsection {* Implementation *}
haftmann@30328
   507
haftmann@30328
   508
datatype 'a seq = Empty | Insert "'a" "'a pred" | Join "'a pred" "'a seq"
haftmann@30328
   509
haftmann@30328
   510
primrec pred_of_seq :: "'a seq \<Rightarrow> 'a pred" where
haftmann@30328
   511
    "pred_of_seq Empty = \<bottom>"
haftmann@30328
   512
  | "pred_of_seq (Insert x P) = single x \<squnion> P"
haftmann@30328
   513
  | "pred_of_seq (Join P xq) = P \<squnion> pred_of_seq xq"
haftmann@30328
   514
haftmann@30328
   515
definition Seq :: "(unit \<Rightarrow> 'a seq) \<Rightarrow> 'a pred" where
haftmann@30328
   516
  "Seq f = pred_of_seq (f ())"
haftmann@30328
   517
haftmann@30328
   518
code_datatype Seq
haftmann@30328
   519
haftmann@30328
   520
primrec member :: "'a seq \<Rightarrow> 'a \<Rightarrow> bool"  where
haftmann@30328
   521
  "member Empty x \<longleftrightarrow> False"
haftmann@30328
   522
  | "member (Insert y P) x \<longleftrightarrow> x = y \<or> eval P x"
haftmann@30328
   523
  | "member (Join P xq) x \<longleftrightarrow> eval P x \<or> member xq x"
haftmann@30328
   524
haftmann@30328
   525
lemma eval_member:
haftmann@30328
   526
  "member xq = eval (pred_of_seq xq)"
haftmann@30328
   527
proof (induct xq)
haftmann@30328
   528
  case Empty show ?case
haftmann@30328
   529
  by (auto simp add: expand_fun_eq elim: botE)
haftmann@30328
   530
next
haftmann@30328
   531
  case Insert show ?case
haftmann@30328
   532
  by (auto simp add: expand_fun_eq elim: supE singleE intro: supI1 supI2 singleI)
haftmann@30328
   533
next
haftmann@30328
   534
  case Join then show ?case
haftmann@30328
   535
  by (auto simp add: expand_fun_eq elim: supE intro: supI1 supI2)
haftmann@30328
   536
qed
haftmann@30328
   537
haftmann@30328
   538
lemma eval_code [code]: "eval (Seq f) = member (f ())"
haftmann@30328
   539
  unfolding Seq_def by (rule sym, rule eval_member)
haftmann@30328
   540
haftmann@30328
   541
lemma single_code [code]:
haftmann@30328
   542
  "single x = Seq (\<lambda>u. Insert x \<bottom>)"
haftmann@30328
   543
  unfolding Seq_def by simp
haftmann@30328
   544
haftmann@30328
   545
primrec "apply" :: "('a \<Rightarrow> 'b Predicate.pred) \<Rightarrow> 'a seq \<Rightarrow> 'b seq" where
haftmann@30328
   546
    "apply f Empty = Empty"
haftmann@30328
   547
  | "apply f (Insert x P) = Join (f x) (Join (P \<guillemotright>= f) Empty)"
haftmann@30328
   548
  | "apply f (Join P xq) = Join (P \<guillemotright>= f) (apply f xq)"
haftmann@30328
   549
haftmann@30328
   550
lemma apply_bind:
haftmann@30328
   551
  "pred_of_seq (apply f xq) = pred_of_seq xq \<guillemotright>= f"
haftmann@30328
   552
proof (induct xq)
haftmann@30328
   553
  case Empty show ?case
haftmann@30328
   554
    by (simp add: bottom_bind)
haftmann@30328
   555
next
haftmann@30328
   556
  case Insert show ?case
haftmann@30328
   557
    by (simp add: single_bind sup_bind)
haftmann@30328
   558
next
haftmann@30328
   559
  case Join then show ?case
haftmann@30328
   560
    by (simp add: sup_bind)
haftmann@30328
   561
qed
haftmann@30328
   562
  
haftmann@30328
   563
lemma bind_code [code]:
haftmann@30328
   564
  "Seq g \<guillemotright>= f = Seq (\<lambda>u. apply f (g ()))"
haftmann@30328
   565
  unfolding Seq_def by (rule sym, rule apply_bind)
haftmann@30328
   566
haftmann@30328
   567
lemma bot_set_code [code]:
haftmann@30328
   568
  "\<bottom> = Seq (\<lambda>u. Empty)"
haftmann@30328
   569
  unfolding Seq_def by simp
haftmann@30328
   570
haftmann@30328
   571
lemma sup_code [code]:
haftmann@30328
   572
  "Seq f \<squnion> Seq g = Seq (\<lambda>u. case f ()
haftmann@30328
   573
    of Empty \<Rightarrow> g ()
haftmann@30328
   574
     | Insert x P \<Rightarrow> Insert x (P \<squnion> Seq g)
haftmann@30328
   575
     | Join P xq \<Rightarrow> Join (Seq g) (Join P xq))" (*FIXME order!?*)
haftmann@30328
   576
proof (cases "f ()")
haftmann@30328
   577
  case Empty
haftmann@30328
   578
  thus ?thesis
haftmann@30328
   579
    unfolding Seq_def by (simp add: sup_commute [of "\<bottom>"] sup_bot)
haftmann@30328
   580
next
haftmann@30328
   581
  case Insert
haftmann@30328
   582
  thus ?thesis
haftmann@30328
   583
    unfolding Seq_def by (simp add: sup_assoc)
haftmann@30328
   584
next
haftmann@30328
   585
  case Join
haftmann@30328
   586
  thus ?thesis
haftmann@30328
   587
    unfolding Seq_def by (simp add: sup_commute [of "pred_of_seq (g ())"] sup_assoc)
haftmann@30328
   588
qed
haftmann@30328
   589
haftmann@30328
   590
haftmann@30328
   591
declare eq_pred_def [code, code del]
haftmann@30328
   592
haftmann@30328
   593
no_notation
haftmann@30328
   594
  inf (infixl "\<sqinter>" 70) and
haftmann@30328
   595
  sup (infixl "\<squnion>" 65) and
haftmann@30328
   596
  Inf ("\<Sqinter>_" [900] 900) and
haftmann@30328
   597
  Sup ("\<Squnion>_" [900] 900) and
haftmann@30328
   598
  top ("\<top>") and
haftmann@30328
   599
  bot ("\<bottom>") and
haftmann@30328
   600
  bind (infixl "\<guillemotright>=" 70)
haftmann@30328
   601
haftmann@30328
   602
hide (open) type pred seq
haftmann@30328
   603
hide (open) const Pred eval single bind if_pred eq_pred not_pred
haftmann@30328
   604
  Empty Insert Join Seq member "apply"
haftmann@30328
   605
haftmann@30328
   606
end