src/HOL/Predicate.thy
author haftmann
Fri, 18 Sep 2009 09:07:48 +0200
changeset 32601 47d0c967c64e
parent 32582 a382876d3290
child 32668 b2de45007537
child 32703 7f9e05b3d0fb
permissions -rw-r--r--
be more cautious wrt. simp rules: sup1_iff, sup2_iff, inf1_iff, inf2_iff, SUP1_iff, SUP2_iff, INF1_iff, INF2_iff are no longer simp by default
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@32601
    78
lemma sup1_iff: "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@32601
    81
lemma sup2_iff: "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)"
haftmann@32601
    85
  by (simp add: sup1_iff 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)"
haftmann@32601
    88
  by (simp add: sup2_iff expand_fun_eq)
berghofe@23741
    89
haftmann@22422
    90
lemma sup1I1 [elim?]: "A x \<Longrightarrow> sup A B x"
haftmann@32601
    91
  by (simp add: sup1_iff)
berghofe@22259
    92
haftmann@22422
    93
lemma sup2I1 [elim?]: "A x y \<Longrightarrow> sup A B x y"
haftmann@32601
    94
  by (simp add: sup2_iff)
berghofe@22259
    95
berghofe@23741
    96
lemma sup1I2 [elim?]: "B x \<Longrightarrow> sup A B x"
haftmann@32601
    97
  by (simp add: sup1_iff)
berghofe@22259
    98
berghofe@23741
    99
lemma sup2I2 [elim?]: "B x y \<Longrightarrow> sup A B x y"
haftmann@32601
   100
  by (simp add: sup2_iff)
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"
haftmann@32601
   108
  by (auto simp add: sup1_iff)
berghofe@22259
   109
haftmann@22422
   110
lemma sup2CI [intro!]: "(~ B x y ==> A x y) ==> sup A B x y"
haftmann@32601
   111
  by (auto simp add: sup2_iff)
berghofe@22259
   112
haftmann@22422
   113
lemma sup1E [elim!]: "sup A B x ==> (A x ==> P) ==> (B x ==> P) ==> P"
haftmann@32601
   114
  by (simp add: sup1_iff) iprover
berghofe@22259
   115
haftmann@22422
   116
lemma sup2E [elim!]: "sup A B x y ==> (A x y ==> P) ==> (B x y ==> P) ==> P"
haftmann@32601
   117
  by (simp add: sup2_iff) iprover
berghofe@22259
   118
berghofe@22259
   119
haftmann@30328
   120
subsubsection {* Binary intersection *}
berghofe@22259
   121
haftmann@32601
   122
lemma inf1_iff: "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@32601
   125
lemma inf2_iff: "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)"
haftmann@32601
   129
  by (simp add: inf1_iff 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)"
haftmann@32601
   132
  by (simp add: inf2_iff expand_fun_eq)
berghofe@23741
   133
haftmann@22422
   134
lemma inf1I [intro!]: "A x ==> B x ==> inf A B x"
haftmann@32601
   135
  by (simp add: inf1_iff)
berghofe@22259
   136
haftmann@22422
   137
lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y"
haftmann@32601
   138
  by (simp add: inf2_iff)
berghofe@22259
   139
haftmann@22422
   140
lemma inf1D1: "inf A B x ==> A x"
haftmann@32601
   141
  by (simp add: inf1_iff)
berghofe@22259
   142
haftmann@22422
   143
lemma inf2D1: "inf A B x y ==> A x y"
haftmann@32601
   144
  by (simp add: inf2_iff)
berghofe@22259
   145
haftmann@22422
   146
lemma inf1D2: "inf A B x ==> B x"
haftmann@32601
   147
  by (simp add: inf1_iff)
berghofe@22259
   148
haftmann@22422
   149
lemma inf2D2: "inf A B x y ==> B x y"
haftmann@32601
   150
  by (simp add: inf2_iff)
berghofe@22259
   151
haftmann@22422
   152
lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P"
haftmann@32601
   153
  by (simp add: inf1_iff)
berghofe@22259
   154
haftmann@22422
   155
lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P"
haftmann@32601
   156
  by (simp add: inf2_iff)
berghofe@22259
   157
berghofe@22259
   158
haftmann@30328
   159
subsubsection {* Unions of families *}
berghofe@22259
   160
haftmann@32601
   161
lemma SUP1_iff: "(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
haftmann@32601
   164
lemma SUP2_iff: "(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"
haftmann@32601
   168
  by (auto simp add: SUP1_iff)
berghofe@22259
   169
berghofe@22430
   170
lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c"
haftmann@32601
   171
  by (auto simp add: SUP2_iff)
berghofe@22259
   172
berghofe@22430
   173
lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R"
haftmann@32601
   174
  by (auto simp add: SUP1_iff)
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"
haftmann@32601
   177
  by (auto simp add: SUP2_iff)
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))"
haftmann@32601
   180
  by (simp add: SUP1_iff 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))"
haftmann@32601
   183
  by (simp add: SUP2_iff expand_fun_eq)
berghofe@22430
   184
berghofe@22430
   185
haftmann@30328
   186
subsubsection {* Intersections of families *}
berghofe@22430
   187
haftmann@32601
   188
lemma INF1_iff: "(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
haftmann@32601
   191
lemma INF2_iff: "(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"
haftmann@32601
   195
  by (auto simp add: INF1_iff)
berghofe@22430
   196
berghofe@22430
   197
lemma INF2_I [intro!]: "(!!x. x : A ==> B x b c) ==> (INF x:A. B x) b c"
haftmann@32601
   198
  by (auto simp add: INF2_iff)
berghofe@22430
   199
berghofe@22430
   200
lemma INF1_D [elim]: "(INF x:A. B x) b ==> a : A ==> B a b"
haftmann@32601
   201
  by (auto simp add: INF1_iff)
berghofe@22430
   202
berghofe@22430
   203
lemma INF2_D [elim]: "(INF x:A. B x) b c ==> a : A ==> B a b c"
haftmann@32601
   204
  by (auto simp add: INF2_iff)
berghofe@22430
   205
berghofe@22430
   206
lemma INF1_E [elim]: "(INF x:A. B x) b ==> (B a b ==> R) ==> (a ~: A ==> R) ==> R"
haftmann@32601
   207
  by (auto simp add: INF1_iff)
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"
haftmann@32601
   210
  by (auto simp add: INF2_iff)
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))"
haftmann@32601
   213
  by (simp add: INF1_iff 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))"
haftmann@32601
   216
  by (simp add: INF2_iff 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
krauss@32231
   224
  pred_comp  :: "['a => 'b => bool, 'b => 'c => bool] => 'a => 'c => bool"
berghofe@22259
   225
    (infixr "OO" 75)
krauss@32231
   226
  for r :: "'a => 'b => bool" and s :: "'b => 'c => bool"
berghofe@22259
   227
where
krauss@32231
   228
  pred_compI [intro]: "r a b ==> s 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@32578
   369
instantiation pred :: (type) "{complete_lattice, boolean_algebra}"
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@31931
   391
  [code del]: "\<Sqinter>A = Pred (INFI A eval)"
haftmann@30328
   392
haftmann@30328
   393
definition
haftmann@31931
   394
  [code del]: "\<Squnion>A = Pred (SUPR A eval)"
haftmann@30328
   395
haftmann@32578
   396
definition
haftmann@32578
   397
  "- P = Pred (- eval P)"
haftmann@32578
   398
haftmann@32578
   399
definition
haftmann@32578
   400
  "P - Q = Pred (eval P - eval Q)"
haftmann@32578
   401
haftmann@32578
   402
instance proof
haftmann@32578
   403
qed (auto simp add: less_eq_pred_def less_pred_def
haftmann@30328
   404
    inf_pred_def sup_pred_def bot_pred_def top_pred_def
haftmann@32578
   405
    Inf_pred_def Sup_pred_def uminus_pred_def minus_pred_def fun_Compl_def bool_Compl_def,
haftmann@30328
   406
    auto simp add: le_fun_def less_fun_def le_bool_def less_bool_def
haftmann@30328
   407
    eval_inject mem_def)
haftmann@30328
   408
berghofe@22259
   409
end
haftmann@30328
   410
haftmann@30328
   411
lemma bind_bind:
haftmann@30328
   412
  "(P \<guillemotright>= Q) \<guillemotright>= R = P \<guillemotright>= (\<lambda>x. Q x \<guillemotright>= R)"
haftmann@30328
   413
  by (auto simp add: bind_def expand_fun_eq)
haftmann@30328
   414
haftmann@30328
   415
lemma bind_single:
haftmann@30328
   416
  "P \<guillemotright>= single = P"
haftmann@30328
   417
  by (simp add: bind_def single_def)
haftmann@30328
   418
haftmann@30328
   419
lemma single_bind:
haftmann@30328
   420
  "single x \<guillemotright>= P = P x"
haftmann@30328
   421
  by (simp add: bind_def single_def)
haftmann@30328
   422
haftmann@30328
   423
lemma bottom_bind:
haftmann@30328
   424
  "\<bottom> \<guillemotright>= P = \<bottom>"
haftmann@30328
   425
  by (auto simp add: bot_pred_def bind_def expand_fun_eq)
haftmann@30328
   426
haftmann@30328
   427
lemma sup_bind:
haftmann@30328
   428
  "(P \<squnion> Q) \<guillemotright>= R = P \<guillemotright>= R \<squnion> Q \<guillemotright>= R"
haftmann@30328
   429
  by (auto simp add: bind_def sup_pred_def expand_fun_eq)
haftmann@30328
   430
haftmann@30328
   431
lemma Sup_bind: "(\<Squnion>A \<guillemotright>= f) = \<Squnion>((\<lambda>x. x \<guillemotright>= f) ` A)"
haftmann@32601
   432
  by (auto simp add: bind_def Sup_pred_def SUP1_iff expand_fun_eq)
haftmann@30328
   433
haftmann@30328
   434
lemma pred_iffI:
haftmann@30328
   435
  assumes "\<And>x. eval A x \<Longrightarrow> eval B x"
haftmann@30328
   436
  and "\<And>x. eval B x \<Longrightarrow> eval A x"
haftmann@30328
   437
  shows "A = B"
haftmann@30328
   438
proof -
haftmann@30328
   439
  from assms have "\<And>x. eval A x \<longleftrightarrow> eval B x" by blast
haftmann@30328
   440
  then show ?thesis by (cases A, cases B) (simp add: expand_fun_eq)
haftmann@30328
   441
qed
haftmann@30328
   442
  
haftmann@30328
   443
lemma singleI: "eval (single x) x"
haftmann@30328
   444
  unfolding single_def by simp
haftmann@30328
   445
haftmann@30328
   446
lemma singleI_unit: "eval (single ()) x"
haftmann@30328
   447
  by simp (rule singleI)
haftmann@30328
   448
haftmann@30328
   449
lemma singleE: "eval (single x) y \<Longrightarrow> (y = x \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   450
  unfolding single_def by simp
haftmann@30328
   451
haftmann@30328
   452
lemma singleE': "eval (single x) y \<Longrightarrow> (x = y \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   453
  by (erule singleE) simp
haftmann@30328
   454
haftmann@30328
   455
lemma bindI: "eval P x \<Longrightarrow> eval (Q x) y \<Longrightarrow> eval (P \<guillemotright>= Q) y"
haftmann@30328
   456
  unfolding bind_def by auto
haftmann@30328
   457
haftmann@30328
   458
lemma bindE: "eval (R \<guillemotright>= Q) y \<Longrightarrow> (\<And>x. eval R x \<Longrightarrow> eval (Q x) y \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   459
  unfolding bind_def by auto
haftmann@30328
   460
haftmann@30328
   461
lemma botE: "eval \<bottom> x \<Longrightarrow> P"
haftmann@30328
   462
  unfolding bot_pred_def by auto
haftmann@30328
   463
haftmann@30328
   464
lemma supI1: "eval A x \<Longrightarrow> eval (A \<squnion> B) x"
haftmann@32601
   465
  unfolding sup_pred_def by (simp add: sup1_iff)
haftmann@30328
   466
haftmann@30328
   467
lemma supI2: "eval B x \<Longrightarrow> eval (A \<squnion> B) x" 
haftmann@32601
   468
  unfolding sup_pred_def by (simp add: sup1_iff)
haftmann@30328
   469
haftmann@30328
   470
lemma supE: "eval (A \<squnion> B) x \<Longrightarrow> (eval A x \<Longrightarrow> P) \<Longrightarrow> (eval B x \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   471
  unfolding sup_pred_def by auto
haftmann@30328
   472
haftmann@32578
   473
lemma single_not_bot [simp]:
haftmann@32578
   474
  "single x \<noteq> \<bottom>"
haftmann@32578
   475
  by (auto simp add: single_def bot_pred_def expand_fun_eq)
haftmann@32578
   476
haftmann@32578
   477
lemma not_bot:
haftmann@32578
   478
  assumes "A \<noteq> \<bottom>"
haftmann@32578
   479
  obtains x where "eval A x"
haftmann@32578
   480
using assms by (cases A)
haftmann@32578
   481
  (auto simp add: bot_pred_def, auto simp add: mem_def)
haftmann@32578
   482
  
haftmann@32578
   483
haftmann@32578
   484
subsubsection {* Emptiness check and definite choice *}
haftmann@32578
   485
haftmann@32578
   486
definition is_empty :: "'a pred \<Rightarrow> bool" where
haftmann@32578
   487
  "is_empty A \<longleftrightarrow> A = \<bottom>"
haftmann@32578
   488
haftmann@32578
   489
lemma is_empty_bot:
haftmann@32578
   490
  "is_empty \<bottom>"
haftmann@32578
   491
  by (simp add: is_empty_def)
haftmann@32578
   492
haftmann@32578
   493
lemma not_is_empty_single:
haftmann@32578
   494
  "\<not> is_empty (single x)"
haftmann@32578
   495
  by (auto simp add: is_empty_def single_def bot_pred_def expand_fun_eq)
haftmann@32578
   496
haftmann@32578
   497
lemma is_empty_sup:
haftmann@32578
   498
  "is_empty (A \<squnion> B) \<longleftrightarrow> is_empty A \<and> is_empty B"
haftmann@32578
   499
  by (auto simp add: is_empty_def intro: sup_eq_bot_eq1 sup_eq_bot_eq2)
haftmann@32578
   500
haftmann@32578
   501
definition singleton :: "'a pred \<Rightarrow> 'a" where
haftmann@32578
   502
  "singleton A = (if \<exists>!x. eval A x then THE x. eval A x else undefined)"
haftmann@32578
   503
haftmann@32578
   504
lemma singleton_eqI:
haftmann@32578
   505
  "\<exists>!x. eval A x \<Longrightarrow> eval A x \<Longrightarrow> singleton A = x"
haftmann@32578
   506
  by (auto simp add: singleton_def)
haftmann@32578
   507
haftmann@32578
   508
lemma eval_singletonI:
haftmann@32578
   509
  "\<exists>!x. eval A x \<Longrightarrow> eval A (singleton A)"
haftmann@32578
   510
proof -
haftmann@32578
   511
  assume assm: "\<exists>!x. eval A x"
haftmann@32578
   512
  then obtain x where "eval A x" ..
haftmann@32578
   513
  moreover with assm have "singleton A = x" by (rule singleton_eqI)
haftmann@32578
   514
  ultimately show ?thesis by simp 
haftmann@32578
   515
qed
haftmann@32578
   516
haftmann@32578
   517
lemma single_singleton:
haftmann@32578
   518
  "\<exists>!x. eval A x \<Longrightarrow> single (singleton A) = A"
haftmann@32578
   519
proof -
haftmann@32578
   520
  assume assm: "\<exists>!x. eval A x"
haftmann@32578
   521
  then have "eval A (singleton A)"
haftmann@32578
   522
    by (rule eval_singletonI)
haftmann@32578
   523
  moreover from assm have "\<And>x. eval A x \<Longrightarrow> singleton A = x"
haftmann@32578
   524
    by (rule singleton_eqI)
haftmann@32578
   525
  ultimately have "eval (single (singleton A)) = eval A"
haftmann@32578
   526
    by (simp (no_asm_use) add: single_def expand_fun_eq) blast
haftmann@32578
   527
  then show ?thesis by (simp add: eval_inject)
haftmann@32578
   528
qed
haftmann@32578
   529
haftmann@32578
   530
lemma singleton_undefinedI:
haftmann@32578
   531
  "\<not> (\<exists>!x. eval A x) \<Longrightarrow> singleton A = undefined"
haftmann@32578
   532
  by (simp add: singleton_def)
haftmann@32578
   533
haftmann@32578
   534
lemma singleton_bot:
haftmann@32578
   535
  "singleton \<bottom> = undefined"
haftmann@32578
   536
  by (auto simp add: bot_pred_def intro: singleton_undefinedI)
haftmann@32578
   537
haftmann@32578
   538
lemma singleton_single:
haftmann@32578
   539
  "singleton (single x) = x"
haftmann@32578
   540
  by (auto simp add: intro: singleton_eqI singleI elim: singleE)
haftmann@32578
   541
haftmann@32578
   542
lemma singleton_sup_single_single:
haftmann@32578
   543
  "singleton (single x \<squnion> single y) = (if x = y then x else undefined)"
haftmann@32578
   544
proof (cases "x = y")
haftmann@32578
   545
  case True then show ?thesis by (simp add: singleton_single)
haftmann@32578
   546
next
haftmann@32578
   547
  case False
haftmann@32578
   548
  have "eval (single x \<squnion> single y) x"
haftmann@32578
   549
    and "eval (single x \<squnion> single y) y"
haftmann@32578
   550
  by (auto intro: supI1 supI2 singleI)
haftmann@32578
   551
  with False have "\<not> (\<exists>!z. eval (single x \<squnion> single y) z)"
haftmann@32578
   552
    by blast
haftmann@32578
   553
  then have "singleton (single x \<squnion> single y) = undefined"
haftmann@32578
   554
    by (rule singleton_undefinedI)
haftmann@32578
   555
  with False show ?thesis by simp
haftmann@32578
   556
qed
haftmann@32578
   557
haftmann@32578
   558
lemma singleton_sup_aux:
haftmann@32578
   559
  "singleton (A \<squnion> B) = (if A = \<bottom> then singleton B
haftmann@32578
   560
    else if B = \<bottom> then singleton A
haftmann@32578
   561
    else singleton
haftmann@32578
   562
      (single (singleton A) \<squnion> single (singleton B)))"
haftmann@32578
   563
proof (cases "(\<exists>!x. eval A x) \<and> (\<exists>!y. eval B y)")
haftmann@32578
   564
  case True then show ?thesis by (simp add: single_singleton)
haftmann@32578
   565
next
haftmann@32578
   566
  case False
haftmann@32578
   567
  from False have A_or_B:
haftmann@32578
   568
    "singleton A = undefined \<or> singleton B = undefined"
haftmann@32578
   569
    by (auto intro!: singleton_undefinedI)
haftmann@32578
   570
  then have rhs: "singleton
haftmann@32578
   571
    (single (singleton A) \<squnion> single (singleton B)) = undefined"
haftmann@32578
   572
    by (auto simp add: singleton_sup_single_single singleton_single)
haftmann@32578
   573
  from False have not_unique:
haftmann@32578
   574
    "\<not> (\<exists>!x. eval A x) \<or> \<not> (\<exists>!y. eval B y)" by simp
haftmann@32578
   575
  show ?thesis proof (cases "A \<noteq> \<bottom> \<and> B \<noteq> \<bottom>")
haftmann@32578
   576
    case True
haftmann@32578
   577
    then obtain a b where a: "eval A a" and b: "eval B b"
haftmann@32578
   578
      by (blast elim: not_bot)
haftmann@32578
   579
    with True not_unique have "\<not> (\<exists>!x. eval (A \<squnion> B) x)"
haftmann@32578
   580
      by (auto simp add: sup_pred_def bot_pred_def)
haftmann@32578
   581
    then have "singleton (A \<squnion> B) = undefined" by (rule singleton_undefinedI)
haftmann@32578
   582
    with True rhs show ?thesis by simp
haftmann@32578
   583
  next
haftmann@32578
   584
    case False then show ?thesis by auto
haftmann@32578
   585
  qed
haftmann@32578
   586
qed
haftmann@32578
   587
haftmann@32578
   588
lemma singleton_sup:
haftmann@32578
   589
  "singleton (A \<squnion> B) = (if A = \<bottom> then singleton B
haftmann@32578
   590
    else if B = \<bottom> then singleton A
haftmann@32578
   591
    else if singleton A = singleton B then singleton A else undefined)"
haftmann@32578
   592
using singleton_sup_aux [of A B] by (simp only: singleton_sup_single_single)
haftmann@32578
   593
haftmann@30328
   594
haftmann@30328
   595
subsubsection {* Derived operations *}
haftmann@30328
   596
haftmann@30328
   597
definition if_pred :: "bool \<Rightarrow> unit pred" where
haftmann@30328
   598
  if_pred_eq: "if_pred b = (if b then single () else \<bottom>)"
haftmann@30328
   599
haftmann@30328
   600
definition not_pred :: "unit pred \<Rightarrow> unit pred" where
haftmann@30328
   601
  not_pred_eq: "not_pred P = (if eval P () then \<bottom> else single ())"
haftmann@30328
   602
haftmann@30328
   603
lemma if_predI: "P \<Longrightarrow> eval (if_pred P) ()"
haftmann@30328
   604
  unfolding if_pred_eq by (auto intro: singleI)
haftmann@30328
   605
haftmann@30328
   606
lemma if_predE: "eval (if_pred b) x \<Longrightarrow> (b \<Longrightarrow> x = () \<Longrightarrow> P) \<Longrightarrow> P"
haftmann@30328
   607
  unfolding if_pred_eq by (cases b) (auto elim: botE)
haftmann@30328
   608
haftmann@30328
   609
lemma not_predI: "\<not> P \<Longrightarrow> eval (not_pred (Pred (\<lambda>u. P))) ()"
haftmann@30328
   610
  unfolding not_pred_eq eval_pred by (auto intro: singleI)
haftmann@30328
   611
haftmann@30328
   612
lemma not_predI': "\<not> eval P () \<Longrightarrow> eval (not_pred P) ()"
haftmann@30328
   613
  unfolding not_pred_eq by (auto intro: singleI)
haftmann@30328
   614
haftmann@30328
   615
lemma not_predE: "eval (not_pred (Pred (\<lambda>u. P))) x \<Longrightarrow> (\<not> P \<Longrightarrow> thesis) \<Longrightarrow> thesis"
haftmann@30328
   616
  unfolding not_pred_eq
haftmann@30328
   617
  by (auto split: split_if_asm elim: botE)
haftmann@30328
   618
haftmann@30328
   619
lemma not_predE': "eval (not_pred P) x \<Longrightarrow> (\<not> eval P x \<Longrightarrow> thesis) \<Longrightarrow> thesis"
haftmann@30328
   620
  unfolding not_pred_eq
haftmann@30328
   621
  by (auto split: split_if_asm elim: botE)
haftmann@30328
   622
haftmann@30328
   623
haftmann@30328
   624
subsubsection {* Implementation *}
haftmann@30328
   625
haftmann@30328
   626
datatype 'a seq = Empty | Insert "'a" "'a pred" | Join "'a pred" "'a seq"
haftmann@30328
   627
haftmann@30328
   628
primrec pred_of_seq :: "'a seq \<Rightarrow> 'a pred" where
haftmann@30328
   629
    "pred_of_seq Empty = \<bottom>"
haftmann@30328
   630
  | "pred_of_seq (Insert x P) = single x \<squnion> P"
haftmann@30328
   631
  | "pred_of_seq (Join P xq) = P \<squnion> pred_of_seq xq"
haftmann@30328
   632
haftmann@30328
   633
definition Seq :: "(unit \<Rightarrow> 'a seq) \<Rightarrow> 'a pred" where
haftmann@30328
   634
  "Seq f = pred_of_seq (f ())"
haftmann@30328
   635
haftmann@30328
   636
code_datatype Seq
haftmann@30328
   637
haftmann@30328
   638
primrec member :: "'a seq \<Rightarrow> 'a \<Rightarrow> bool"  where
haftmann@30328
   639
  "member Empty x \<longleftrightarrow> False"
haftmann@30328
   640
  | "member (Insert y P) x \<longleftrightarrow> x = y \<or> eval P x"
haftmann@30328
   641
  | "member (Join P xq) x \<longleftrightarrow> eval P x \<or> member xq x"
haftmann@30328
   642
haftmann@30328
   643
lemma eval_member:
haftmann@30328
   644
  "member xq = eval (pred_of_seq xq)"
haftmann@30328
   645
proof (induct xq)
haftmann@30328
   646
  case Empty show ?case
haftmann@30328
   647
  by (auto simp add: expand_fun_eq elim: botE)
haftmann@30328
   648
next
haftmann@30328
   649
  case Insert show ?case
haftmann@30328
   650
  by (auto simp add: expand_fun_eq elim: supE singleE intro: supI1 supI2 singleI)
haftmann@30328
   651
next
haftmann@30328
   652
  case Join then show ?case
haftmann@30328
   653
  by (auto simp add: expand_fun_eq elim: supE intro: supI1 supI2)
haftmann@30328
   654
qed
haftmann@30328
   655
haftmann@30328
   656
lemma eval_code [code]: "eval (Seq f) = member (f ())"
haftmann@30328
   657
  unfolding Seq_def by (rule sym, rule eval_member)
haftmann@30328
   658
haftmann@30328
   659
lemma single_code [code]:
haftmann@30328
   660
  "single x = Seq (\<lambda>u. Insert x \<bottom>)"
haftmann@30328
   661
  unfolding Seq_def by simp
haftmann@30328
   662
haftmann@30328
   663
primrec "apply" :: "('a \<Rightarrow> 'b Predicate.pred) \<Rightarrow> 'a seq \<Rightarrow> 'b seq" where
haftmann@30328
   664
    "apply f Empty = Empty"
haftmann@30328
   665
  | "apply f (Insert x P) = Join (f x) (Join (P \<guillemotright>= f) Empty)"
haftmann@30328
   666
  | "apply f (Join P xq) = Join (P \<guillemotright>= f) (apply f xq)"
haftmann@30328
   667
haftmann@30328
   668
lemma apply_bind:
haftmann@30328
   669
  "pred_of_seq (apply f xq) = pred_of_seq xq \<guillemotright>= f"
haftmann@30328
   670
proof (induct xq)
haftmann@30328
   671
  case Empty show ?case
haftmann@30328
   672
    by (simp add: bottom_bind)
haftmann@30328
   673
next
haftmann@30328
   674
  case Insert show ?case
haftmann@30328
   675
    by (simp add: single_bind sup_bind)
haftmann@30328
   676
next
haftmann@30328
   677
  case Join then show ?case
haftmann@30328
   678
    by (simp add: sup_bind)
haftmann@30328
   679
qed
haftmann@30328
   680
  
haftmann@30328
   681
lemma bind_code [code]:
haftmann@30328
   682
  "Seq g \<guillemotright>= f = Seq (\<lambda>u. apply f (g ()))"
haftmann@30328
   683
  unfolding Seq_def by (rule sym, rule apply_bind)
haftmann@30328
   684
haftmann@30328
   685
lemma bot_set_code [code]:
haftmann@30328
   686
  "\<bottom> = Seq (\<lambda>u. Empty)"
haftmann@30328
   687
  unfolding Seq_def by simp
haftmann@30328
   688
haftmann@30376
   689
primrec adjunct :: "'a pred \<Rightarrow> 'a seq \<Rightarrow> 'a seq" where
haftmann@30376
   690
    "adjunct P Empty = Join P Empty"
haftmann@30376
   691
  | "adjunct P (Insert x Q) = Insert x (Q \<squnion> P)"
haftmann@30376
   692
  | "adjunct P (Join Q xq) = Join Q (adjunct P xq)"
haftmann@30376
   693
haftmann@30376
   694
lemma adjunct_sup:
haftmann@30376
   695
  "pred_of_seq (adjunct P xq) = P \<squnion> pred_of_seq xq"
haftmann@30376
   696
  by (induct xq) (simp_all add: sup_assoc sup_commute sup_left_commute)
haftmann@30376
   697
haftmann@30328
   698
lemma sup_code [code]:
haftmann@30328
   699
  "Seq f \<squnion> Seq g = Seq (\<lambda>u. case f ()
haftmann@30328
   700
    of Empty \<Rightarrow> g ()
haftmann@30328
   701
     | Insert x P \<Rightarrow> Insert x (P \<squnion> Seq g)
haftmann@30376
   702
     | Join P xq \<Rightarrow> adjunct (Seq g) (Join P xq))"
haftmann@30328
   703
proof (cases "f ()")
haftmann@30328
   704
  case Empty
haftmann@30328
   705
  thus ?thesis
haftmann@30376
   706
    unfolding Seq_def by (simp add: sup_commute [of "\<bottom>"]  sup_bot)
haftmann@30328
   707
next
haftmann@30328
   708
  case Insert
haftmann@30328
   709
  thus ?thesis
haftmann@30328
   710
    unfolding Seq_def by (simp add: sup_assoc)
haftmann@30328
   711
next
haftmann@30328
   712
  case Join
haftmann@30328
   713
  thus ?thesis
haftmann@30376
   714
    unfolding Seq_def
haftmann@30376
   715
    by (simp add: adjunct_sup sup_assoc sup_commute sup_left_commute)
haftmann@30328
   716
qed
haftmann@30328
   717
haftmann@30430
   718
primrec contained :: "'a seq \<Rightarrow> 'a pred \<Rightarrow> bool" where
haftmann@30430
   719
    "contained Empty Q \<longleftrightarrow> True"
haftmann@30430
   720
  | "contained (Insert x P) Q \<longleftrightarrow> eval Q x \<and> P \<le> Q"
haftmann@30430
   721
  | "contained (Join P xq) Q \<longleftrightarrow> P \<le> Q \<and> contained xq Q"
haftmann@30430
   722
haftmann@30430
   723
lemma single_less_eq_eval:
haftmann@30430
   724
  "single x \<le> P \<longleftrightarrow> eval P x"
haftmann@30430
   725
  by (auto simp add: single_def less_eq_pred_def mem_def)
haftmann@30430
   726
haftmann@30430
   727
lemma contained_less_eq:
haftmann@30430
   728
  "contained xq Q \<longleftrightarrow> pred_of_seq xq \<le> Q"
haftmann@30430
   729
  by (induct xq) (simp_all add: single_less_eq_eval)
haftmann@30430
   730
haftmann@30430
   731
lemma less_eq_pred_code [code]:
haftmann@30430
   732
  "Seq f \<le> Q = (case f ()
haftmann@30430
   733
   of Empty \<Rightarrow> True
haftmann@30430
   734
    | Insert x P \<Rightarrow> eval Q x \<and> P \<le> Q
haftmann@30430
   735
    | Join P xq \<Rightarrow> P \<le> Q \<and> contained xq Q)"
haftmann@30430
   736
  by (cases "f ()")
haftmann@30430
   737
    (simp_all add: Seq_def single_less_eq_eval contained_less_eq)
haftmann@30430
   738
haftmann@30430
   739
lemma eq_pred_code [code]:
haftmann@31133
   740
  fixes P Q :: "'a pred"
haftmann@30430
   741
  shows "eq_class.eq P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P"
haftmann@30430
   742
  unfolding eq by auto
haftmann@30430
   743
haftmann@30430
   744
lemma [code]:
haftmann@30430
   745
  "pred_case f P = f (eval P)"
haftmann@30430
   746
  by (cases P) simp
haftmann@30430
   747
haftmann@30430
   748
lemma [code]:
haftmann@30430
   749
  "pred_rec f P = f (eval P)"
haftmann@30430
   750
  by (cases P) simp
haftmann@30328
   751
bulwahn@31105
   752
inductive eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where "eq x x"
bulwahn@31105
   753
bulwahn@31105
   754
lemma eq_is_eq: "eq x y \<equiv> (x = y)"
haftmann@31108
   755
  by (rule eq_reflection) (auto intro: eq.intros elim: eq.cases)
haftmann@30948
   756
haftmann@31216
   757
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a pred \<Rightarrow> 'b pred" where
haftmann@31216
   758
  "map f P = P \<guillemotright>= (single o f)"
haftmann@31216
   759
haftmann@32578
   760
primrec null :: "'a seq \<Rightarrow> bool" where
haftmann@32578
   761
    "null Empty \<longleftrightarrow> True"
haftmann@32578
   762
  | "null (Insert x P) \<longleftrightarrow> False"
haftmann@32578
   763
  | "null (Join P xq) \<longleftrightarrow> is_empty P \<and> null xq"
haftmann@32578
   764
haftmann@32578
   765
lemma null_is_empty:
haftmann@32578
   766
  "null xq \<longleftrightarrow> is_empty (pred_of_seq xq)"
haftmann@32578
   767
  by (induct xq) (simp_all add: is_empty_bot not_is_empty_single is_empty_sup)
haftmann@32578
   768
haftmann@32578
   769
lemma is_empty_code [code]:
haftmann@32578
   770
  "is_empty (Seq f) \<longleftrightarrow> null (f ())"
haftmann@32578
   771
  by (simp add: null_is_empty Seq_def)
haftmann@32578
   772
haftmann@32578
   773
primrec the_only :: "'a seq \<Rightarrow> 'a" where
haftmann@32578
   774
  [code del]: "the_only Empty = undefined"
haftmann@32578
   775
  | "the_only (Insert x P) = (if is_empty P then x else let y = singleton P in if x = y then x else undefined)"
haftmann@32578
   776
  | "the_only (Join P xq) = (if is_empty P then the_only xq else if null xq then singleton P
haftmann@32578
   777
       else let x = singleton P; y = the_only xq in
haftmann@32578
   778
       if x = y then x else undefined)"
haftmann@32578
   779
haftmann@32578
   780
lemma the_only_singleton:
haftmann@32578
   781
  "the_only xq = singleton (pred_of_seq xq)"
haftmann@32578
   782
  by (induct xq)
haftmann@32578
   783
    (auto simp add: singleton_bot singleton_single is_empty_def
haftmann@32578
   784
    null_is_empty Let_def singleton_sup)
haftmann@32578
   785
haftmann@32578
   786
lemma singleton_code [code]:
haftmann@32578
   787
  "singleton (Seq f) = (case f ()
haftmann@32578
   788
   of Empty \<Rightarrow> undefined
haftmann@32578
   789
    | Insert x P \<Rightarrow> if is_empty P then x
haftmann@32578
   790
        else let y = singleton P in
haftmann@32578
   791
          if x = y then x else undefined
haftmann@32578
   792
    | Join P xq \<Rightarrow> if is_empty P then the_only xq
haftmann@32578
   793
        else if null xq then singleton P
haftmann@32578
   794
        else let x = singleton P; y = the_only xq in
haftmann@32578
   795
          if x = y then x else undefined)"
haftmann@32578
   796
  by (cases "f ()")
haftmann@32578
   797
   (auto simp add: Seq_def the_only_singleton is_empty_def
haftmann@32578
   798
      null_is_empty singleton_bot singleton_single singleton_sup Let_def)
haftmann@32578
   799
haftmann@30948
   800
ML {*
haftmann@30948
   801
signature PREDICATE =
haftmann@30948
   802
sig
haftmann@30948
   803
  datatype 'a pred = Seq of (unit -> 'a seq)
haftmann@30948
   804
  and 'a seq = Empty | Insert of 'a * 'a pred | Join of 'a pred * 'a seq
haftmann@30959
   805
  val yield: 'a pred -> ('a * 'a pred) option
haftmann@30959
   806
  val yieldn: int -> 'a pred -> 'a list * 'a pred
haftmann@31222
   807
  val map: ('a -> 'b) -> 'a pred -> 'b pred
haftmann@30948
   808
end;
haftmann@30948
   809
haftmann@30948
   810
structure Predicate : PREDICATE =
haftmann@30948
   811
struct
haftmann@30948
   812
haftmann@30959
   813
@{code_datatype pred = Seq};
haftmann@30959
   814
@{code_datatype seq = Empty | Insert | Join};
haftmann@30959
   815
haftmann@32372
   816
fun yield (@{code Seq} f) = next (f ())
haftmann@30959
   817
and next @{code Empty} = NONE
haftmann@30959
   818
  | next (@{code Insert} (x, P)) = SOME (x, P)
haftmann@30959
   819
  | next (@{code Join} (P, xq)) = (case yield P
haftmann@30959
   820
     of NONE => next xq
haftmann@30959
   821
      | SOME (x, Q) => SOME (x, @{code Seq} (fn _ => @{code Join} (Q, xq))))
haftmann@30959
   822
haftmann@30959
   823
fun anamorph f k x = (if k = 0 then ([], x)
haftmann@30959
   824
  else case f x
haftmann@30959
   825
   of NONE => ([], x)
haftmann@30959
   826
    | SOME (v, y) => let
haftmann@30959
   827
        val (vs, z) = anamorph f (k - 1) y
haftmann@30959
   828
      in (v :: vs, z) end)
haftmann@30959
   829
haftmann@30959
   830
fun yieldn P = anamorph yield P;
haftmann@30948
   831
haftmann@31222
   832
fun map f = @{code map} f;
haftmann@31222
   833
haftmann@30948
   834
end;
haftmann@30948
   835
*}
haftmann@30948
   836
haftmann@30948
   837
code_reserved Eval Predicate
haftmann@30948
   838
haftmann@30948
   839
code_type pred and seq
haftmann@30948
   840
  (Eval "_/ Predicate.pred" and "_/ Predicate.seq")
haftmann@30948
   841
haftmann@30948
   842
code_const Seq and Empty and Insert and Join
haftmann@30948
   843
  (Eval "Predicate.Seq" and "Predicate.Empty" and "Predicate.Insert/ (_,/ _)" and "Predicate.Join/ (_,/ _)")
haftmann@30948
   844
haftmann@31122
   845
text {* dummy setup for @{text code_pred} and @{text values} keywords *}
haftmann@31108
   846
haftmann@31108
   847
ML {*
haftmann@31122
   848
local
haftmann@31122
   849
haftmann@31122
   850
structure P = OuterParse;
haftmann@31122
   851
haftmann@31122
   852
val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
haftmann@31122
   853
haftmann@31122
   854
in
haftmann@31122
   855
haftmann@31122
   856
val _ = OuterSyntax.local_theory_to_proof "code_pred" "sets up goal for cases rule from given introduction rules and compiles predicate"
haftmann@31122
   857
  OuterKeyword.thy_goal (P.term_group >> (K (Proof.theorem_i NONE (K I) [[]])));
haftmann@31122
   858
haftmann@31216
   859
val _ = OuterSyntax.improper_command "values" "enumerate and print comprehensions"
haftmann@31122
   860
  OuterKeyword.diag ((opt_modes -- P.term)
haftmann@31122
   861
    >> (fn (modes, t) => Toplevel.no_timing o Toplevel.keep
haftmann@31122
   862
        (K ())));
haftmann@31122
   863
haftmann@31122
   864
end
haftmann@31108
   865
*}
haftmann@30959
   866
haftmann@30328
   867
no_notation
haftmann@30328
   868
  inf (infixl "\<sqinter>" 70) and
haftmann@30328
   869
  sup (infixl "\<squnion>" 65) and
haftmann@30328
   870
  Inf ("\<Sqinter>_" [900] 900) and
haftmann@30328
   871
  Sup ("\<Squnion>_" [900] 900) and
haftmann@30328
   872
  top ("\<top>") and
haftmann@30328
   873
  bot ("\<bottom>") and
haftmann@30328
   874
  bind (infixl "\<guillemotright>=" 70)
haftmann@30328
   875
haftmann@30328
   876
hide (open) type pred seq
haftmann@32582
   877
hide (open) const Pred eval single bind is_empty singleton if_pred not_pred
haftmann@32582
   878
  Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map
haftmann@30328
   879
haftmann@30328
   880
end