src/HOL/Predicate.thy
author haftmann
Fri, 29 Jul 2011 19:47:55 +0200
changeset 44878 b5e7594061ce
parent 41798 efa734d9b221
child 44897 d5e28a49e16e
permissions -rw-r--r--
tuned proofs
     1 (*  Title:      HOL/Predicate.thy
     2     Author:     Stefan Berghofer and Lukas Bulwahn and Florian Haftmann, TU Muenchen
     3 *)
     4 
     5 header {* Predicates as relations and enumerations *}
     6 
     7 theory Predicate
     8 imports Inductive Relation
     9 begin
    10 
    11 notation
    12   bot ("\<bottom>") and
    13   top ("\<top>") and
    14   inf (infixl "\<sqinter>" 70) and
    15   sup (infixl "\<squnion>" 65) and
    16   Inf ("\<Sqinter>_" [900] 900) and
    17   Sup ("\<Squnion>_" [900] 900)
    18 
    19 syntax (xsymbols)
    20   "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
    21   "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
    22   "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
    23   "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
    24 
    25 
    26 subsection {* Predicates as (complete) lattices *}
    27 
    28 
    29 text {*
    30   Handy introduction and elimination rules for @{text "\<le>"}
    31   on unary and binary predicates
    32 *}
    33 
    34 lemma predicate1I:
    35   assumes PQ: "\<And>x. P x \<Longrightarrow> Q x"
    36   shows "P \<le> Q"
    37   apply (rule le_funI)
    38   apply (rule le_boolI)
    39   apply (rule PQ)
    40   apply assumption
    41   done
    42 
    43 lemma predicate1D [Pure.dest?, dest?]:
    44   "P \<le> Q \<Longrightarrow> P x \<Longrightarrow> Q x"
    45   apply (erule le_funE)
    46   apply (erule le_boolE)
    47   apply assumption+
    48   done
    49 
    50 lemma rev_predicate1D:
    51   "P x ==> P <= Q ==> Q x"
    52   by (rule predicate1D)
    53 
    54 lemma predicate2I [Pure.intro!, intro!]:
    55   assumes PQ: "\<And>x y. P x y \<Longrightarrow> Q x y"
    56   shows "P \<le> Q"
    57   apply (rule le_funI)+
    58   apply (rule le_boolI)
    59   apply (rule PQ)
    60   apply assumption
    61   done
    62 
    63 lemma predicate2D [Pure.dest, dest]:
    64   "P \<le> Q \<Longrightarrow> P x y \<Longrightarrow> Q x y"
    65   apply (erule le_funE)+
    66   apply (erule le_boolE)
    67   apply assumption+
    68   done
    69 
    70 lemma rev_predicate2D:
    71   "P x y ==> P <= Q ==> Q x y"
    72   by (rule predicate2D)
    73 
    74 
    75 subsubsection {* Equality *}
    76 
    77 lemma pred_equals_eq: "((\<lambda>x. x \<in> R) = (\<lambda>x. x \<in> S)) = (R = S)"
    78   by (simp add: mem_def)
    79 
    80 lemma pred_equals_eq2 [pred_set_conv]: "((\<lambda>x y. (x, y) \<in> R) = (\<lambda>x y. (x, y) \<in> S)) = (R = S)"
    81   by (simp add: fun_eq_iff mem_def)
    82 
    83 
    84 subsubsection {* Order relation *}
    85 
    86 lemma pred_subset_eq: "((\<lambda>x. x \<in> R) <= (\<lambda>x. x \<in> S)) = (R <= S)"
    87   by (simp add: mem_def)
    88 
    89 lemma pred_subset_eq2 [pred_set_conv]: "((\<lambda>x y. (x, y) \<in> R) <= (\<lambda>x y. (x, y) \<in> S)) = (R <= S)"
    90   by fast
    91 
    92 
    93 subsubsection {* Top and bottom elements *}
    94 
    95 lemma bot1E [no_atp, elim!]: "bot x \<Longrightarrow> P"
    96   by (simp add: bot_fun_def)
    97 
    98 lemma bot2E [elim!]: "bot x y \<Longrightarrow> P"
    99   by (simp add: bot_fun_def)
   100 
   101 lemma bot_empty_eq: "bot = (\<lambda>x. x \<in> {})"
   102   by (auto simp add: fun_eq_iff)
   103 
   104 lemma bot_empty_eq2: "bot = (\<lambda>x y. (x, y) \<in> {})"
   105   by (auto simp add: fun_eq_iff)
   106 
   107 lemma top1I [intro!]: "top x"
   108   by (simp add: top_fun_def)
   109 
   110 lemma top2I [intro!]: "top x y"
   111   by (simp add: top_fun_def)
   112 
   113 
   114 subsubsection {* Binary intersection *}
   115 
   116 lemma inf1I [intro!]: "A x ==> B x ==> inf A B x"
   117   by (simp add: inf_fun_def)
   118 
   119 lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y"
   120   by (simp add: inf_fun_def)
   121 
   122 lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P"
   123   by (simp add: inf_fun_def)
   124 
   125 lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P"
   126   by (simp add: inf_fun_def)
   127 
   128 lemma inf1D1: "inf A B x ==> A x"
   129   by (simp add: inf_fun_def)
   130 
   131 lemma inf2D1: "inf A B x y ==> A x y"
   132   by (simp add: inf_fun_def)
   133 
   134 lemma inf1D2: "inf A B x ==> B x"
   135   by (simp add: inf_fun_def)
   136 
   137 lemma inf2D2: "inf A B x y ==> B x y"
   138   by (simp add: inf_fun_def)
   139 
   140 lemma inf_Int_eq: "inf (\<lambda>x. x \<in> R) (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<inter> S)"
   141   by (simp add: inf_fun_def mem_def)
   142 
   143 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)"
   144   by (simp add: inf_fun_def mem_def)
   145 
   146 
   147 subsubsection {* Binary union *}
   148 
   149 lemma sup1I1 [elim?]: "A x \<Longrightarrow> sup A B x"
   150   by (simp add: sup_fun_def)
   151 
   152 lemma sup2I1 [elim?]: "A x y \<Longrightarrow> sup A B x y"
   153   by (simp add: sup_fun_def)
   154 
   155 lemma sup1I2 [elim?]: "B x \<Longrightarrow> sup A B x"
   156   by (simp add: sup_fun_def)
   157 
   158 lemma sup2I2 [elim?]: "B x y \<Longrightarrow> sup A B x y"
   159   by (simp add: sup_fun_def)
   160 
   161 lemma sup1E [elim!]: "sup A B x ==> (A x ==> P) ==> (B x ==> P) ==> P"
   162   by (simp add: sup_fun_def) iprover
   163 
   164 lemma sup2E [elim!]: "sup A B x y ==> (A x y ==> P) ==> (B x y ==> P) ==> P"
   165   by (simp add: sup_fun_def) iprover
   166 
   167 text {*
   168   \medskip Classical introduction rule: no commitment to @{text A} vs
   169   @{text B}.
   170 *}
   171 
   172 lemma sup1CI [intro!]: "(~ B x ==> A x) ==> sup A B x"
   173   by (auto simp add: sup_fun_def)
   174 
   175 lemma sup2CI [intro!]: "(~ B x y ==> A x y) ==> sup A B x y"
   176   by (auto simp add: sup_fun_def)
   177 
   178 lemma sup_Un_eq: "sup (\<lambda>x. x \<in> R) (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<union> S)"
   179   by (simp add: sup_fun_def mem_def)
   180 
   181 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)"
   182   by (simp add: sup_fun_def mem_def)
   183 
   184 
   185 subsubsection {* Intersections of families *}
   186 
   187 lemma INF1_iff: "(INF x:A. B x) b = (ALL x:A. B x b)"
   188   by (simp add: INFI_apply)
   189 
   190 lemma INF2_iff: "(INF x:A. B x) b c = (ALL x:A. B x b c)"
   191   by (simp add: INFI_apply)
   192 
   193 lemma INF1_I [intro!]: "(!!x. x : A ==> B x b) ==> (INF x:A. B x) b"
   194   by (auto simp add: INFI_apply)
   195 
   196 lemma INF2_I [intro!]: "(!!x. x : A ==> B x b c) ==> (INF x:A. B x) b c"
   197   by (auto simp add: INFI_apply)
   198 
   199 lemma INF1_D [elim]: "(INF x:A. B x) b ==> a : A ==> B a b"
   200   by (auto simp add: INFI_apply)
   201 
   202 lemma INF2_D [elim]: "(INF x:A. B x) b c ==> a : A ==> B a b c"
   203   by (auto simp add: INFI_apply)
   204 
   205 lemma INF1_E [elim]: "(INF x:A. B x) b ==> (B a b ==> R) ==> (a ~: A ==> R) ==> R"
   206   by (auto simp add: INFI_apply)
   207 
   208 lemma INF2_E [elim]: "(INF x:A. B x) b c ==> (B a b c ==> R) ==> (a ~: A ==> R) ==> R"
   209   by (auto simp add: INFI_apply)
   210 
   211 lemma INF_INT_eq: "(INF i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (INT i. r i))"
   212   by (simp add: INFI_apply fun_eq_iff)
   213 
   214 lemma INF_INT_eq2: "(INF i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (INT i. r i))"
   215   by (simp add: INFI_apply fun_eq_iff)
   216 
   217 
   218 subsubsection {* Unions of families *}
   219 
   220 lemma SUP1_iff: "(SUP x:A. B x) b = (EX x:A. B x b)"
   221   by (simp add: SUPR_apply)
   222 
   223 lemma SUP2_iff: "(SUP x:A. B x) b c = (EX x:A. B x b c)"
   224   by (simp add: SUPR_apply)
   225 
   226 lemma SUP1_I [intro]: "a : A ==> B a b ==> (SUP x:A. B x) b"
   227   by (auto simp add: SUPR_apply)
   228 
   229 lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c"
   230   by (auto simp add: SUPR_apply)
   231 
   232 lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R"
   233   by (auto simp add: SUPR_apply)
   234 
   235 lemma SUP2_E [elim!]: "(SUP x:A. B x) b c ==> (!!x. x : A ==> B x b c ==> R) ==> R"
   236   by (auto simp add: SUPR_apply)
   237 
   238 lemma SUP_UN_eq: "(SUP i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (UN i. r i))"
   239   by (simp add: SUPR_apply fun_eq_iff)
   240 
   241 lemma SUP_UN_eq2: "(SUP i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (UN i. r i))"
   242   by (simp add: SUPR_apply fun_eq_iff)
   243 
   244 
   245 subsection {* Predicates as relations *}
   246 
   247 subsubsection {* Composition  *}
   248 
   249 inductive
   250   pred_comp  :: "['a => 'b => bool, 'b => 'c => bool] => 'a => 'c => bool"
   251     (infixr "OO" 75)
   252   for r :: "'a => 'b => bool" and s :: "'b => 'c => bool"
   253 where
   254   pred_compI [intro]: "r a b ==> s b c ==> (r OO s) a c"
   255 
   256 inductive_cases pred_compE [elim!]: "(r OO s) a c"
   257 
   258 lemma pred_comp_rel_comp_eq [pred_set_conv]:
   259   "((\<lambda>x y. (x, y) \<in> r) OO (\<lambda>x y. (x, y) \<in> s)) = (\<lambda>x y. (x, y) \<in> r O s)"
   260   by (auto simp add: fun_eq_iff)
   261 
   262 
   263 subsubsection {* Converse *}
   264 
   265 inductive
   266   conversep :: "('a => 'b => bool) => 'b => 'a => bool"
   267     ("(_^--1)" [1000] 1000)
   268   for r :: "'a => 'b => bool"
   269 where
   270   conversepI: "r a b ==> r^--1 b a"
   271 
   272 notation (xsymbols)
   273   conversep  ("(_\<inverse>\<inverse>)" [1000] 1000)
   274 
   275 lemma conversepD:
   276   assumes ab: "r^--1 a b"
   277   shows "r b a" using ab
   278   by cases simp
   279 
   280 lemma conversep_iff [iff]: "r^--1 a b = r b a"
   281   by (iprover intro: conversepI dest: conversepD)
   282 
   283 lemma conversep_converse_eq [pred_set_conv]:
   284   "(\<lambda>x y. (x, y) \<in> r)^--1 = (\<lambda>x y. (x, y) \<in> r^-1)"
   285   by (auto simp add: fun_eq_iff)
   286 
   287 lemma conversep_conversep [simp]: "(r^--1)^--1 = r"
   288   by (iprover intro: order_antisym conversepI dest: conversepD)
   289 
   290 lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1"
   291   by (iprover intro: order_antisym conversepI pred_compI
   292     elim: pred_compE dest: conversepD)
   293 
   294 lemma converse_meet: "(inf r s)^--1 = inf r^--1 s^--1"
   295   by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD)
   296 
   297 lemma converse_join: "(sup r s)^--1 = sup r^--1 s^--1"
   298   by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD)
   299 
   300 lemma conversep_noteq [simp]: "(op ~=)^--1 = op ~="
   301   by (auto simp add: fun_eq_iff)
   302 
   303 lemma conversep_eq [simp]: "(op =)^--1 = op ="
   304   by (auto simp add: fun_eq_iff)
   305 
   306 
   307 subsubsection {* Domain *}
   308 
   309 inductive
   310   DomainP :: "('a => 'b => bool) => 'a => bool"
   311   for r :: "'a => 'b => bool"
   312 where
   313   DomainPI [intro]: "r a b ==> DomainP r a"
   314 
   315 inductive_cases DomainPE [elim!]: "DomainP r a"
   316 
   317 lemma DomainP_Domain_eq [pred_set_conv]: "DomainP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Domain r)"
   318   by (blast intro!: Orderings.order_antisym predicate1I)
   319 
   320 
   321 subsubsection {* Range *}
   322 
   323 inductive
   324   RangeP :: "('a => 'b => bool) => 'b => bool"
   325   for r :: "'a => 'b => bool"
   326 where
   327   RangePI [intro]: "r a b ==> RangeP r b"
   328 
   329 inductive_cases RangePE [elim!]: "RangeP r b"
   330 
   331 lemma RangeP_Range_eq [pred_set_conv]: "RangeP (\<lambda>x y. (x, y) \<in> r) = (\<lambda>x. x \<in> Range r)"
   332   by (blast intro!: Orderings.order_antisym predicate1I)
   333 
   334 
   335 subsubsection {* Inverse image *}
   336 
   337 definition
   338   inv_imagep :: "('b => 'b => bool) => ('a => 'b) => 'a => 'a => bool" where
   339   "inv_imagep r f == %x y. r (f x) (f y)"
   340 
   341 lemma [pred_set_conv]: "inv_imagep (\<lambda>x y. (x, y) \<in> r) f = (\<lambda>x y. (x, y) \<in> inv_image r f)"
   342   by (simp add: inv_image_def inv_imagep_def)
   343 
   344 lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)"
   345   by (simp add: inv_imagep_def)
   346 
   347 
   348 subsubsection {* Powerset *}
   349 
   350 definition Powp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool" where
   351   "Powp A == \<lambda>B. \<forall>x \<in> B. A x"
   352 
   353 lemma Powp_Pow_eq [pred_set_conv]: "Powp (\<lambda>x. x \<in> A) = (\<lambda>x. x \<in> Pow A)"
   354   by (auto simp add: Powp_def fun_eq_iff)
   355 
   356 lemmas Powp_mono [mono] = Pow_mono [to_pred pred_subset_eq]
   357 
   358 
   359 subsubsection {* Properties of relations *}
   360 
   361 abbreviation antisymP :: "('a => 'a => bool) => bool" where
   362   "antisymP r == antisym {(x, y). r x y}"
   363 
   364 abbreviation transP :: "('a => 'a => bool) => bool" where
   365   "transP r == trans {(x, y). r x y}"
   366 
   367 abbreviation single_valuedP :: "('a => 'b => bool) => bool" where
   368   "single_valuedP r == single_valued {(x, y). r x y}"
   369 
   370 (*FIXME inconsistencies: abbreviations vs. definitions, suffix `P` vs. suffix `p`*)
   371 
   372 definition reflp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   373   "reflp r \<longleftrightarrow> refl {(x, y). r x y}"
   374 
   375 definition symp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   376   "symp r \<longleftrightarrow> sym {(x, y). r x y}"
   377 
   378 definition transp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   379   "transp r \<longleftrightarrow> trans {(x, y). r x y}"
   380 
   381 lemma reflpI:
   382   "(\<And>x. r x x) \<Longrightarrow> reflp r"
   383   by (auto intro: refl_onI simp add: reflp_def)
   384 
   385 lemma reflpE:
   386   assumes "reflp r"
   387   obtains "r x x"
   388   using assms by (auto dest: refl_onD simp add: reflp_def)
   389 
   390 lemma sympI:
   391   "(\<And>x y. r x y \<Longrightarrow> r y x) \<Longrightarrow> symp r"
   392   by (auto intro: symI simp add: symp_def)
   393 
   394 lemma sympE:
   395   assumes "symp r" and "r x y"
   396   obtains "r y x"
   397   using assms by (auto dest: symD simp add: symp_def)
   398 
   399 lemma transpI:
   400   "(\<And>x y z. r x y \<Longrightarrow> r y z \<Longrightarrow> r x z) \<Longrightarrow> transp r"
   401   by (auto intro: transI simp add: transp_def)
   402   
   403 lemma transpE:
   404   assumes "transp r" and "r x y" and "r y z"
   405   obtains "r x z"
   406   using assms by (auto dest: transD simp add: transp_def)
   407 
   408 
   409 subsection {* Predicates as enumerations *}
   410 
   411 subsubsection {* The type of predicate enumerations (a monad) *}
   412 
   413 datatype 'a pred = Pred "'a \<Rightarrow> bool"
   414 
   415 primrec eval :: "'a pred \<Rightarrow> 'a \<Rightarrow> bool" where
   416   eval_pred: "eval (Pred f) = f"
   417 
   418 lemma Pred_eval [simp]:
   419   "Pred (eval x) = x"
   420   by (cases x) simp
   421 
   422 lemma pred_eqI:
   423   "(\<And>w. eval P w \<longleftrightarrow> eval Q w) \<Longrightarrow> P = Q"
   424   by (cases P, cases Q) (auto simp add: fun_eq_iff)
   425 
   426 lemma eval_mem [simp]:
   427   "x \<in> eval P \<longleftrightarrow> eval P x"
   428   by (simp add: mem_def)
   429 
   430 lemma eq_mem [simp]:
   431   "x \<in> (op =) y \<longleftrightarrow> x = y"
   432   by (auto simp add: mem_def)
   433 
   434 instantiation pred :: (type) "{complete_lattice, boolean_algebra}"
   435 begin
   436 
   437 definition
   438   "P \<le> Q \<longleftrightarrow> eval P \<le> eval Q"
   439 
   440 definition
   441   "P < Q \<longleftrightarrow> eval P < eval Q"
   442 
   443 definition
   444   "\<bottom> = Pred \<bottom>"
   445 
   446 lemma eval_bot [simp]:
   447   "eval \<bottom>  = \<bottom>"
   448   by (simp add: bot_pred_def)
   449 
   450 definition
   451   "\<top> = Pred \<top>"
   452 
   453 lemma eval_top [simp]:
   454   "eval \<top>  = \<top>"
   455   by (simp add: top_pred_def)
   456 
   457 definition
   458   "P \<sqinter> Q = Pred (eval P \<sqinter> eval Q)"
   459 
   460 lemma eval_inf [simp]:
   461   "eval (P \<sqinter> Q) = eval P \<sqinter> eval Q"
   462   by (simp add: inf_pred_def)
   463 
   464 definition
   465   "P \<squnion> Q = Pred (eval P \<squnion> eval Q)"
   466 
   467 lemma eval_sup [simp]:
   468   "eval (P \<squnion> Q) = eval P \<squnion> eval Q"
   469   by (simp add: sup_pred_def)
   470 
   471 definition
   472   "\<Sqinter>A = Pred (INFI A eval)"
   473 
   474 lemma eval_Inf [simp]:
   475   "eval (\<Sqinter>A) = INFI A eval"
   476   by (simp add: Inf_pred_def)
   477 
   478 definition
   479   "\<Squnion>A = Pred (SUPR A eval)"
   480 
   481 lemma eval_Sup [simp]:
   482   "eval (\<Squnion>A) = SUPR A eval"
   483   by (simp add: Sup_pred_def)
   484 
   485 definition
   486   "- P = Pred (- eval P)"
   487 
   488 lemma eval_compl [simp]:
   489   "eval (- P) = - eval P"
   490   by (simp add: uminus_pred_def)
   491 
   492 definition
   493   "P - Q = Pred (eval P - eval Q)"
   494 
   495 lemma eval_minus [simp]:
   496   "eval (P - Q) = eval P - eval Q"
   497   by (simp add: minus_pred_def)
   498 
   499 instance proof
   500 qed (auto intro!: pred_eqI simp add: less_eq_pred_def less_pred_def uminus_apply minus_apply)
   501 
   502 end
   503 
   504 lemma eval_INFI [simp]:
   505   "eval (INFI A f) = INFI A (eval \<circ> f)"
   506   by (unfold INFI_def) simp
   507 
   508 lemma eval_SUPR [simp]:
   509   "eval (SUPR A f) = SUPR A (eval \<circ> f)"
   510   by (unfold SUPR_def) simp
   511 
   512 definition single :: "'a \<Rightarrow> 'a pred" where
   513   "single x = Pred ((op =) x)"
   514 
   515 lemma eval_single [simp]:
   516   "eval (single x) = (op =) x"
   517   by (simp add: single_def)
   518 
   519 definition bind :: "'a pred \<Rightarrow> ('a \<Rightarrow> 'b pred) \<Rightarrow> 'b pred" (infixl "\<guillemotright>=" 70) where
   520   "P \<guillemotright>= f = (SUPR {x. eval P x} f)"
   521 
   522 lemma eval_bind [simp]:
   523   "eval (P \<guillemotright>= f) = eval (SUPR {x. eval P x} f)"
   524   by (simp add: bind_def)
   525 
   526 lemma bind_bind:
   527   "(P \<guillemotright>= Q) \<guillemotright>= R = P \<guillemotright>= (\<lambda>x. Q x \<guillemotright>= R)"
   528   by (rule pred_eqI) auto
   529 
   530 lemma bind_single:
   531   "P \<guillemotright>= single = P"
   532   by (rule pred_eqI) auto
   533 
   534 lemma single_bind:
   535   "single x \<guillemotright>= P = P x"
   536   by (rule pred_eqI) auto
   537 
   538 lemma bottom_bind:
   539   "\<bottom> \<guillemotright>= P = \<bottom>"
   540   by (rule pred_eqI) auto
   541 
   542 lemma sup_bind:
   543   "(P \<squnion> Q) \<guillemotright>= R = P \<guillemotright>= R \<squnion> Q \<guillemotright>= R"
   544   by (rule pred_eqI) auto
   545 
   546 lemma Sup_bind:
   547   "(\<Squnion>A \<guillemotright>= f) = \<Squnion>((\<lambda>x. x \<guillemotright>= f) ` A)"
   548   by (rule pred_eqI) auto
   549 
   550 lemma pred_iffI:
   551   assumes "\<And>x. eval A x \<Longrightarrow> eval B x"
   552   and "\<And>x. eval B x \<Longrightarrow> eval A x"
   553   shows "A = B"
   554   using assms by (auto intro: pred_eqI)
   555   
   556 lemma singleI: "eval (single x) x"
   557   by simp
   558 
   559 lemma singleI_unit: "eval (single ()) x"
   560   by simp
   561 
   562 lemma singleE: "eval (single x) y \<Longrightarrow> (y = x \<Longrightarrow> P) \<Longrightarrow> P"
   563   by simp
   564 
   565 lemma singleE': "eval (single x) y \<Longrightarrow> (x = y \<Longrightarrow> P) \<Longrightarrow> P"
   566   by simp
   567 
   568 lemma bindI: "eval P x \<Longrightarrow> eval (Q x) y \<Longrightarrow> eval (P \<guillemotright>= Q) y"
   569   by auto
   570 
   571 lemma bindE: "eval (R \<guillemotright>= Q) y \<Longrightarrow> (\<And>x. eval R x \<Longrightarrow> eval (Q x) y \<Longrightarrow> P) \<Longrightarrow> P"
   572   by auto
   573 
   574 lemma botE: "eval \<bottom> x \<Longrightarrow> P"
   575   by auto
   576 
   577 lemma supI1: "eval A x \<Longrightarrow> eval (A \<squnion> B) x"
   578   by auto
   579 
   580 lemma supI2: "eval B x \<Longrightarrow> eval (A \<squnion> B) x" 
   581   by auto
   582 
   583 lemma supE: "eval (A \<squnion> B) x \<Longrightarrow> (eval A x \<Longrightarrow> P) \<Longrightarrow> (eval B x \<Longrightarrow> P) \<Longrightarrow> P"
   584   by auto
   585 
   586 lemma single_not_bot [simp]:
   587   "single x \<noteq> \<bottom>"
   588   by (auto simp add: single_def bot_pred_def fun_eq_iff)
   589 
   590 lemma not_bot:
   591   assumes "A \<noteq> \<bottom>"
   592   obtains x where "eval A x"
   593   using assms by (cases A)
   594     (auto simp add: bot_pred_def, auto simp add: mem_def)
   595   
   596 
   597 subsubsection {* Emptiness check and definite choice *}
   598 
   599 definition is_empty :: "'a pred \<Rightarrow> bool" where
   600   "is_empty A \<longleftrightarrow> A = \<bottom>"
   601 
   602 lemma is_empty_bot:
   603   "is_empty \<bottom>"
   604   by (simp add: is_empty_def)
   605 
   606 lemma not_is_empty_single:
   607   "\<not> is_empty (single x)"
   608   by (auto simp add: is_empty_def single_def bot_pred_def fun_eq_iff)
   609 
   610 lemma is_empty_sup:
   611   "is_empty (A \<squnion> B) \<longleftrightarrow> is_empty A \<and> is_empty B"
   612   by (auto simp add: is_empty_def)
   613 
   614 definition singleton :: "(unit \<Rightarrow> 'a) \<Rightarrow> 'a pred \<Rightarrow> 'a" where
   615   "singleton dfault A = (if \<exists>!x. eval A x then THE x. eval A x else dfault ())"
   616 
   617 lemma singleton_eqI:
   618   "\<exists>!x. eval A x \<Longrightarrow> eval A x \<Longrightarrow> singleton dfault A = x"
   619   by (auto simp add: singleton_def)
   620 
   621 lemma eval_singletonI:
   622   "\<exists>!x. eval A x \<Longrightarrow> eval A (singleton dfault A)"
   623 proof -
   624   assume assm: "\<exists>!x. eval A x"
   625   then obtain x where "eval A x" ..
   626   moreover with assm have "singleton dfault A = x" by (rule singleton_eqI)
   627   ultimately show ?thesis by simp 
   628 qed
   629 
   630 lemma single_singleton:
   631   "\<exists>!x. eval A x \<Longrightarrow> single (singleton dfault A) = A"
   632 proof -
   633   assume assm: "\<exists>!x. eval A x"
   634   then have "eval A (singleton dfault A)"
   635     by (rule eval_singletonI)
   636   moreover from assm have "\<And>x. eval A x \<Longrightarrow> singleton dfault A = x"
   637     by (rule singleton_eqI)
   638   ultimately have "eval (single (singleton dfault A)) = eval A"
   639     by (simp (no_asm_use) add: single_def fun_eq_iff) blast
   640   then have "\<And>x. eval (single (singleton dfault A)) x = eval A x"
   641     by simp
   642   then show ?thesis by (rule pred_eqI)
   643 qed
   644 
   645 lemma singleton_undefinedI:
   646   "\<not> (\<exists>!x. eval A x) \<Longrightarrow> singleton dfault A = dfault ()"
   647   by (simp add: singleton_def)
   648 
   649 lemma singleton_bot:
   650   "singleton dfault \<bottom> = dfault ()"
   651   by (auto simp add: bot_pred_def intro: singleton_undefinedI)
   652 
   653 lemma singleton_single:
   654   "singleton dfault (single x) = x"
   655   by (auto simp add: intro: singleton_eqI singleI elim: singleE)
   656 
   657 lemma singleton_sup_single_single:
   658   "singleton dfault (single x \<squnion> single y) = (if x = y then x else dfault ())"
   659 proof (cases "x = y")
   660   case True then show ?thesis by (simp add: singleton_single)
   661 next
   662   case False
   663   have "eval (single x \<squnion> single y) x"
   664     and "eval (single x \<squnion> single y) y"
   665   by (auto intro: supI1 supI2 singleI)
   666   with False have "\<not> (\<exists>!z. eval (single x \<squnion> single y) z)"
   667     by blast
   668   then have "singleton dfault (single x \<squnion> single y) = dfault ()"
   669     by (rule singleton_undefinedI)
   670   with False show ?thesis by simp
   671 qed
   672 
   673 lemma singleton_sup_aux:
   674   "singleton dfault (A \<squnion> B) = (if A = \<bottom> then singleton dfault B
   675     else if B = \<bottom> then singleton dfault A
   676     else singleton dfault
   677       (single (singleton dfault A) \<squnion> single (singleton dfault B)))"
   678 proof (cases "(\<exists>!x. eval A x) \<and> (\<exists>!y. eval B y)")
   679   case True then show ?thesis by (simp add: single_singleton)
   680 next
   681   case False
   682   from False have A_or_B:
   683     "singleton dfault A = dfault () \<or> singleton dfault B = dfault ()"
   684     by (auto intro!: singleton_undefinedI)
   685   then have rhs: "singleton dfault
   686     (single (singleton dfault A) \<squnion> single (singleton dfault B)) = dfault ()"
   687     by (auto simp add: singleton_sup_single_single singleton_single)
   688   from False have not_unique:
   689     "\<not> (\<exists>!x. eval A x) \<or> \<not> (\<exists>!y. eval B y)" by simp
   690   show ?thesis proof (cases "A \<noteq> \<bottom> \<and> B \<noteq> \<bottom>")
   691     case True
   692     then obtain a b where a: "eval A a" and b: "eval B b"
   693       by (blast elim: not_bot)
   694     with True not_unique have "\<not> (\<exists>!x. eval (A \<squnion> B) x)"
   695       by (auto simp add: sup_pred_def bot_pred_def)
   696     then have "singleton dfault (A \<squnion> B) = dfault ()" by (rule singleton_undefinedI)
   697     with True rhs show ?thesis by simp
   698   next
   699     case False then show ?thesis by auto
   700   qed
   701 qed
   702 
   703 lemma singleton_sup:
   704   "singleton dfault (A \<squnion> B) = (if A = \<bottom> then singleton dfault B
   705     else if B = \<bottom> then singleton dfault A
   706     else if singleton dfault A = singleton dfault B then singleton dfault A else dfault ())"
   707 using singleton_sup_aux [of dfault A B] by (simp only: singleton_sup_single_single)
   708 
   709 
   710 subsubsection {* Derived operations *}
   711 
   712 definition if_pred :: "bool \<Rightarrow> unit pred" where
   713   if_pred_eq: "if_pred b = (if b then single () else \<bottom>)"
   714 
   715 definition holds :: "unit pred \<Rightarrow> bool" where
   716   holds_eq: "holds P = eval P ()"
   717 
   718 definition not_pred :: "unit pred \<Rightarrow> unit pred" where
   719   not_pred_eq: "not_pred P = (if eval P () then \<bottom> else single ())"
   720 
   721 lemma if_predI: "P \<Longrightarrow> eval (if_pred P) ()"
   722   unfolding if_pred_eq by (auto intro: singleI)
   723 
   724 lemma if_predE: "eval (if_pred b) x \<Longrightarrow> (b \<Longrightarrow> x = () \<Longrightarrow> P) \<Longrightarrow> P"
   725   unfolding if_pred_eq by (cases b) (auto elim: botE)
   726 
   727 lemma not_predI: "\<not> P \<Longrightarrow> eval (not_pred (Pred (\<lambda>u. P))) ()"
   728   unfolding not_pred_eq eval_pred by (auto intro: singleI)
   729 
   730 lemma not_predI': "\<not> eval P () \<Longrightarrow> eval (not_pred P) ()"
   731   unfolding not_pred_eq by (auto intro: singleI)
   732 
   733 lemma not_predE: "eval (not_pred (Pred (\<lambda>u. P))) x \<Longrightarrow> (\<not> P \<Longrightarrow> thesis) \<Longrightarrow> thesis"
   734   unfolding not_pred_eq
   735   by (auto split: split_if_asm elim: botE)
   736 
   737 lemma not_predE': "eval (not_pred P) x \<Longrightarrow> (\<not> eval P x \<Longrightarrow> thesis) \<Longrightarrow> thesis"
   738   unfolding not_pred_eq
   739   by (auto split: split_if_asm elim: botE)
   740 lemma "f () = False \<or> f () = True"
   741 by simp
   742 
   743 lemma closure_of_bool_cases [no_atp]:
   744   fixes f :: "unit \<Rightarrow> bool"
   745   assumes "f = (\<lambda>u. False) \<Longrightarrow> P f"
   746   assumes "f = (\<lambda>u. True) \<Longrightarrow> P f"
   747   shows "P f"
   748 proof -
   749   have "f = (\<lambda>u. False) \<or> f = (\<lambda>u. True)"
   750     apply (cases "f ()")
   751     apply (rule disjI2)
   752     apply (rule ext)
   753     apply (simp add: unit_eq)
   754     apply (rule disjI1)
   755     apply (rule ext)
   756     apply (simp add: unit_eq)
   757     done
   758   from this assms show ?thesis by blast
   759 qed
   760 
   761 lemma unit_pred_cases:
   762   assumes "P \<bottom>"
   763   assumes "P (single ())"
   764   shows "P Q"
   765 using assms unfolding bot_pred_def Collect_def empty_def single_def proof (cases Q)
   766   fix f
   767   assume "P (Pred (\<lambda>u. False))" "P (Pred (\<lambda>u. () = u))"
   768   then have "P (Pred f)" 
   769     by (cases _ f rule: closure_of_bool_cases) simp_all
   770   moreover assume "Q = Pred f"
   771   ultimately show "P Q" by simp
   772 qed
   773   
   774 lemma holds_if_pred:
   775   "holds (if_pred b) = b"
   776 unfolding if_pred_eq holds_eq
   777 by (cases b) (auto intro: singleI elim: botE)
   778 
   779 lemma if_pred_holds:
   780   "if_pred (holds P) = P"
   781 unfolding if_pred_eq holds_eq
   782 by (rule unit_pred_cases) (auto intro: singleI elim: botE)
   783 
   784 lemma is_empty_holds:
   785   "is_empty P \<longleftrightarrow> \<not> holds P"
   786 unfolding is_empty_def holds_eq
   787 by (rule unit_pred_cases) (auto elim: botE intro: singleI)
   788 
   789 definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a pred \<Rightarrow> 'b pred" where
   790   "map f P = P \<guillemotright>= (single o f)"
   791 
   792 lemma eval_map [simp]:
   793   "eval (map f P) = image f (eval P)"
   794   by (auto simp add: map_def)
   795 
   796 enriched_type map: map
   797   by (auto intro!: pred_eqI simp add: fun_eq_iff image_compose)
   798 
   799 
   800 subsubsection {* Implementation *}
   801 
   802 datatype 'a seq = Empty | Insert "'a" "'a pred" | Join "'a pred" "'a seq"
   803 
   804 primrec pred_of_seq :: "'a seq \<Rightarrow> 'a pred" where
   805     "pred_of_seq Empty = \<bottom>"
   806   | "pred_of_seq (Insert x P) = single x \<squnion> P"
   807   | "pred_of_seq (Join P xq) = P \<squnion> pred_of_seq xq"
   808 
   809 definition Seq :: "(unit \<Rightarrow> 'a seq) \<Rightarrow> 'a pred" where
   810   "Seq f = pred_of_seq (f ())"
   811 
   812 code_datatype Seq
   813 
   814 primrec member :: "'a seq \<Rightarrow> 'a \<Rightarrow> bool"  where
   815   "member Empty x \<longleftrightarrow> False"
   816   | "member (Insert y P) x \<longleftrightarrow> x = y \<or> eval P x"
   817   | "member (Join P xq) x \<longleftrightarrow> eval P x \<or> member xq x"
   818 
   819 lemma eval_member:
   820   "member xq = eval (pred_of_seq xq)"
   821 proof (induct xq)
   822   case Empty show ?case
   823   by (auto simp add: fun_eq_iff elim: botE)
   824 next
   825   case Insert show ?case
   826   by (auto simp add: fun_eq_iff elim: supE singleE intro: supI1 supI2 singleI)
   827 next
   828   case Join then show ?case
   829   by (auto simp add: fun_eq_iff elim: supE intro: supI1 supI2)
   830 qed
   831 
   832 lemma eval_code [code]: "eval (Seq f) = member (f ())"
   833   unfolding Seq_def by (rule sym, rule eval_member)
   834 
   835 lemma single_code [code]:
   836   "single x = Seq (\<lambda>u. Insert x \<bottom>)"
   837   unfolding Seq_def by simp
   838 
   839 primrec "apply" :: "('a \<Rightarrow> 'b pred) \<Rightarrow> 'a seq \<Rightarrow> 'b seq" where
   840     "apply f Empty = Empty"
   841   | "apply f (Insert x P) = Join (f x) (Join (P \<guillemotright>= f) Empty)"
   842   | "apply f (Join P xq) = Join (P \<guillemotright>= f) (apply f xq)"
   843 
   844 lemma apply_bind:
   845   "pred_of_seq (apply f xq) = pred_of_seq xq \<guillemotright>= f"
   846 proof (induct xq)
   847   case Empty show ?case
   848     by (simp add: bottom_bind)
   849 next
   850   case Insert show ?case
   851     by (simp add: single_bind sup_bind)
   852 next
   853   case Join then show ?case
   854     by (simp add: sup_bind)
   855 qed
   856   
   857 lemma bind_code [code]:
   858   "Seq g \<guillemotright>= f = Seq (\<lambda>u. apply f (g ()))"
   859   unfolding Seq_def by (rule sym, rule apply_bind)
   860 
   861 lemma bot_set_code [code]:
   862   "\<bottom> = Seq (\<lambda>u. Empty)"
   863   unfolding Seq_def by simp
   864 
   865 primrec adjunct :: "'a pred \<Rightarrow> 'a seq \<Rightarrow> 'a seq" where
   866     "adjunct P Empty = Join P Empty"
   867   | "adjunct P (Insert x Q) = Insert x (Q \<squnion> P)"
   868   | "adjunct P (Join Q xq) = Join Q (adjunct P xq)"
   869 
   870 lemma adjunct_sup:
   871   "pred_of_seq (adjunct P xq) = P \<squnion> pred_of_seq xq"
   872   by (induct xq) (simp_all add: sup_assoc sup_commute sup_left_commute)
   873 
   874 lemma sup_code [code]:
   875   "Seq f \<squnion> Seq g = Seq (\<lambda>u. case f ()
   876     of Empty \<Rightarrow> g ()
   877      | Insert x P \<Rightarrow> Insert x (P \<squnion> Seq g)
   878      | Join P xq \<Rightarrow> adjunct (Seq g) (Join P xq))"
   879 proof (cases "f ()")
   880   case Empty
   881   thus ?thesis
   882     unfolding Seq_def by (simp add: sup_commute [of "\<bottom>"])
   883 next
   884   case Insert
   885   thus ?thesis
   886     unfolding Seq_def by (simp add: sup_assoc)
   887 next
   888   case Join
   889   thus ?thesis
   890     unfolding Seq_def
   891     by (simp add: adjunct_sup sup_assoc sup_commute sup_left_commute)
   892 qed
   893 
   894 primrec contained :: "'a seq \<Rightarrow> 'a pred \<Rightarrow> bool" where
   895     "contained Empty Q \<longleftrightarrow> True"
   896   | "contained (Insert x P) Q \<longleftrightarrow> eval Q x \<and> P \<le> Q"
   897   | "contained (Join P xq) Q \<longleftrightarrow> P \<le> Q \<and> contained xq Q"
   898 
   899 lemma single_less_eq_eval:
   900   "single x \<le> P \<longleftrightarrow> eval P x"
   901   by (auto simp add: single_def less_eq_pred_def mem_def)
   902 
   903 lemma contained_less_eq:
   904   "contained xq Q \<longleftrightarrow> pred_of_seq xq \<le> Q"
   905   by (induct xq) (simp_all add: single_less_eq_eval)
   906 
   907 lemma less_eq_pred_code [code]:
   908   "Seq f \<le> Q = (case f ()
   909    of Empty \<Rightarrow> True
   910     | Insert x P \<Rightarrow> eval Q x \<and> P \<le> Q
   911     | Join P xq \<Rightarrow> P \<le> Q \<and> contained xq Q)"
   912   by (cases "f ()")
   913     (simp_all add: Seq_def single_less_eq_eval contained_less_eq)
   914 
   915 lemma eq_pred_code [code]:
   916   fixes P Q :: "'a pred"
   917   shows "HOL.equal P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P"
   918   by (auto simp add: equal)
   919 
   920 lemma [code nbe]:
   921   "HOL.equal (x :: 'a pred) x \<longleftrightarrow> True"
   922   by (fact equal_refl)
   923 
   924 lemma [code]:
   925   "pred_case f P = f (eval P)"
   926   by (cases P) simp
   927 
   928 lemma [code]:
   929   "pred_rec f P = f (eval P)"
   930   by (cases P) simp
   931 
   932 inductive eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where "eq x x"
   933 
   934 lemma eq_is_eq: "eq x y \<equiv> (x = y)"
   935   by (rule eq_reflection) (auto intro: eq.intros elim: eq.cases)
   936 
   937 primrec null :: "'a seq \<Rightarrow> bool" where
   938     "null Empty \<longleftrightarrow> True"
   939   | "null (Insert x P) \<longleftrightarrow> False"
   940   | "null (Join P xq) \<longleftrightarrow> is_empty P \<and> null xq"
   941 
   942 lemma null_is_empty:
   943   "null xq \<longleftrightarrow> is_empty (pred_of_seq xq)"
   944   by (induct xq) (simp_all add: is_empty_bot not_is_empty_single is_empty_sup)
   945 
   946 lemma is_empty_code [code]:
   947   "is_empty (Seq f) \<longleftrightarrow> null (f ())"
   948   by (simp add: null_is_empty Seq_def)
   949 
   950 primrec the_only :: "(unit \<Rightarrow> 'a) \<Rightarrow> 'a seq \<Rightarrow> 'a" where
   951   [code del]: "the_only dfault Empty = dfault ()"
   952   | "the_only dfault (Insert x P) = (if is_empty P then x else let y = singleton dfault P in if x = y then x else dfault ())"
   953   | "the_only dfault (Join P xq) = (if is_empty P then the_only dfault xq else if null xq then singleton dfault P
   954        else let x = singleton dfault P; y = the_only dfault xq in
   955        if x = y then x else dfault ())"
   956 
   957 lemma the_only_singleton:
   958   "the_only dfault xq = singleton dfault (pred_of_seq xq)"
   959   by (induct xq)
   960     (auto simp add: singleton_bot singleton_single is_empty_def
   961     null_is_empty Let_def singleton_sup)
   962 
   963 lemma singleton_code [code]:
   964   "singleton dfault (Seq f) = (case f ()
   965    of Empty \<Rightarrow> dfault ()
   966     | Insert x P \<Rightarrow> if is_empty P then x
   967         else let y = singleton dfault P in
   968           if x = y then x else dfault ()
   969     | Join P xq \<Rightarrow> if is_empty P then the_only dfault xq
   970         else if null xq then singleton dfault P
   971         else let x = singleton dfault P; y = the_only dfault xq in
   972           if x = y then x else dfault ())"
   973   by (cases "f ()")
   974    (auto simp add: Seq_def the_only_singleton is_empty_def
   975       null_is_empty singleton_bot singleton_single singleton_sup Let_def)
   976 
   977 definition not_unique :: "'a pred => 'a"
   978 where
   979   [code del]: "not_unique A = (THE x. eval A x)"
   980 
   981 definition the :: "'a pred => 'a"
   982 where
   983   "the A = (THE x. eval A x)"
   984 
   985 lemma the_eqI:
   986   "(THE x. eval P x) = x \<Longrightarrow> the P = x"
   987   by (simp add: the_def)
   988 
   989 lemma the_eq [code]: "the A = singleton (\<lambda>x. not_unique A) A"
   990   by (rule the_eqI) (simp add: singleton_def not_unique_def)
   991 
   992 code_abort not_unique
   993 
   994 code_reflect Predicate
   995   datatypes pred = Seq and seq = Empty | Insert | Join
   996   functions map
   997 
   998 ML {*
   999 signature PREDICATE =
  1000 sig
  1001   datatype 'a pred = Seq of (unit -> 'a seq)
  1002   and 'a seq = Empty | Insert of 'a * 'a pred | Join of 'a pred * 'a seq
  1003   val yield: 'a pred -> ('a * 'a pred) option
  1004   val yieldn: int -> 'a pred -> 'a list * 'a pred
  1005   val map: ('a -> 'b) -> 'a pred -> 'b pred
  1006 end;
  1007 
  1008 structure Predicate : PREDICATE =
  1009 struct
  1010 
  1011 datatype pred = datatype Predicate.pred
  1012 datatype seq = datatype Predicate.seq
  1013 
  1014 fun map f = Predicate.map f;
  1015 
  1016 fun yield (Seq f) = next (f ())
  1017 and next Empty = NONE
  1018   | next (Insert (x, P)) = SOME (x, P)
  1019   | next (Join (P, xq)) = (case yield P
  1020      of NONE => next xq
  1021       | SOME (x, Q) => SOME (x, Seq (fn _ => Join (Q, xq))));
  1022 
  1023 fun anamorph f k x = (if k = 0 then ([], x)
  1024   else case f x
  1025    of NONE => ([], x)
  1026     | SOME (v, y) => let
  1027         val (vs, z) = anamorph f (k - 1) y
  1028       in (v :: vs, z) end);
  1029 
  1030 fun yieldn P = anamorph yield P;
  1031 
  1032 end;
  1033 *}
  1034 
  1035 no_notation
  1036   bot ("\<bottom>") and
  1037   top ("\<top>") and
  1038   inf (infixl "\<sqinter>" 70) and
  1039   sup (infixl "\<squnion>" 65) and
  1040   Inf ("\<Sqinter>_" [900] 900) and
  1041   Sup ("\<Squnion>_" [900] 900) and
  1042   bind (infixl "\<guillemotright>=" 70)
  1043 
  1044 no_syntax (xsymbols)
  1045   "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
  1046   "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
  1047   "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
  1048   "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
  1049 
  1050 hide_type (open) pred seq
  1051 hide_const (open) Pred eval single bind is_empty singleton if_pred not_pred holds
  1052   Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map not_unique the
  1053 
  1054 end