Nominal/activities/tphols09/IDW/Conversions.thy
changeset 415 f1be8028a4a9
equal deleted inserted replaced
414:05e5d68c9627 415:f1be8028a4a9
       
     1 theory Conversions
       
     2 imports Main
       
     3 begin
       
     4 
       
     5 section {* Basic conversions *}
       
     6 
       
     7 ML {* Conv.all_conv : cterm -> thm *}
       
     8 ML {* Conv.all_conv *}
       
     9 
       
    10 text {* Always succeeds *}
       
    11 
       
    12 ML {* Conv.all_conv @{cterm "42::int"} *}
       
    13 
       
    14 text {* Always fails *}
       
    15 
       
    16 (*
       
    17 ML {* Conv.no_conv @{cterm "42::int"} *}
       
    18 *)
       
    19 
       
    20 text {* Rewrite with a single rule *}
       
    21 
       
    22 ML {*
       
    23 val rev_Cons = mk_meta_eq @{thm rev.simps(2)}
       
    24 *}
       
    25 
       
    26 ML {*
       
    27 Conv.rewr_conv rev_Cons @{cterm "rev [1::int, 2]"}
       
    28 *}
       
    29 
       
    30 
       
    31 
       
    32 section {* Combining conversions (``conversionals'') *}
       
    33 
       
    34 text {* Sequencing *}
       
    35 
       
    36 ML {*
       
    37 val add_0_right = mk_meta_eq @{thm monoid_add_class.add_0_right}
       
    38 *}
       
    39 
       
    40 ML {*
       
    41 val mult_1_left = mk_meta_eq @{thm monoid_mult_class.mult_1_left}
       
    42 *}
       
    43 
       
    44 ML {*
       
    45 (Conv.rewr_conv add_0_right then_conv Conv.rewr_conv mult_1_left)
       
    46 @{cterm "((1::int) * x) + 0"}
       
    47 *}
       
    48 
       
    49 text {* Alternative *}
       
    50 
       
    51 ML {*
       
    52 val rev_Nil = mk_meta_eq @{thm rev.simps(1)}
       
    53 *}
       
    54 
       
    55 ML {*
       
    56 (Conv.rewr_conv rev_Nil else_conv Conv.rewr_conv rev_Cons)
       
    57 @{cterm "rev [1::int, 2, 3]"}
       
    58 *}
       
    59 
       
    60 text {* Try conversion (yields reflexivity instead of exception) *}
       
    61 
       
    62 ML {*
       
    63 Conv.try_conv (Conv.rewr_conv rev_Nil) @{cterm "[1::int, 2]"}
       
    64 *}
       
    65 
       
    66 text {* Descend into subterms *}
       
    67 
       
    68 ML {*
       
    69 Conv.combination_conv
       
    70   (Conv.combination_conv
       
    71     Conv.all_conv
       
    72     (Conv.rewr_conv rev_Cons))
       
    73   (Conv.rewr_conv rev_Cons)
       
    74   @{cterm "rev [1::int, 2] @ rev [3, 4]"}
       
    75 *}
       
    76 
       
    77 ML {*
       
    78 Conv.combination_conv
       
    79   (Conv.arg_conv
       
    80     (Conv.rewr_conv rev_Cons))
       
    81   (Conv.rewr_conv rev_Cons)
       
    82   @{cterm "rev [1::int, 2] @ rev [3, 4]"}
       
    83 *}
       
    84 
       
    85 ML {*
       
    86 Conv.abs_conv (fn (v, ctxt) =>
       
    87   Conv.abs_conv (fn (v', ctxt') =>
       
    88     Conv.rewr_conv rev_Cons) ctxt)
       
    89   @{context}
       
    90   @{cterm "\<lambda>x y. rev [x, y]"}
       
    91 *}
       
    92 
       
    93 
       
    94 
       
    95 section {* Simple bottom-up rewriting, using Isabelle's conversion library *}
       
    96 
       
    97 text {* Descend into immediate subterms *}
       
    98 
       
    99 ML {*
       
   100 fun subc conv ctxt = 
       
   101   Conv.comb_conv (conv ctxt) else_conv
       
   102   Conv.abs_conv (conv o snd) ctxt else_conv
       
   103   Conv.all_conv;
       
   104 *}
       
   105 
       
   106 text {* The ct argument is necessary to avoid nontermination! *}
       
   107 
       
   108 ML {*
       
   109 fun botc conv ctxt ct =
       
   110   (subc (botc conv) ctxt then_conv
       
   111    Conv.try_conv (conv then_conv botc conv ctxt)) ct
       
   112 *}
       
   113 
       
   114 text {* Running example: reversing lists *}
       
   115 
       
   116 ML {*
       
   117 val eqns = map mk_meta_eq (@{thms "append.simps"} @ @{thms "rev.simps"});
       
   118 *}
       
   119 
       
   120 ML {*
       
   121 val rev_int = @{cterm "rev :: int list \<Rightarrow> int list"};
       
   122 *}
       
   123 
       
   124 text {* Produce lists of length i *}
       
   125 
       
   126 ML {*
       
   127 fun mk_upto thy i = Thm.cterm_of thy (HOLogic.mk_list HOLogic.intT
       
   128   (map (HOLogic.mk_number HOLogic.intT) (1 upto i)));
       
   129 *}
       
   130 
       
   131 ML {*
       
   132 val ct = Thm.capply rev_int (mk_upto @{theory} 100)
       
   133 *}
       
   134 
       
   135 text {* Fully rewrite the term (slow, i.e. 40 secs on my laptop) *}
       
   136 
       
   137 ML {*
       
   138 timeit (fn () =>
       
   139   botc (Conv.first_conv (map Conv.rewr_conv eqns)) @{context} ct)
       
   140 *}
       
   141 
       
   142 text {* Also rewrites inside quantifiers, thanks to abs_conv *}
       
   143 
       
   144 ML {*
       
   145 botc (Conv.first_conv (map Conv.rewr_conv eqns)) @{context}
       
   146   @{cterm "\<forall>x y z. P (rev [x, y, z])"}
       
   147 *}
       
   148 
       
   149 
       
   150 
       
   151 section {* Using exceptions for signalling unchanged terms *}
       
   152 
       
   153 text {*
       
   154   Motivation: avoid unnecessary applications of reflexivity
       
   155 *}
       
   156 
       
   157 ML {*
       
   158 infix 1 then_conv';
       
   159 infix 0 else_conv';
       
   160 
       
   161 exception Fail;
       
   162 exception Unchanged;
       
   163 
       
   164 fun (cv1 else_conv' cv2) ct =
       
   165   cv1 ct handle Fail => cv2 ct;
       
   166 
       
   167 fun all_conv' ct = raise Unchanged;
       
   168 
       
   169 fun no_conv' ct = raise Fail;
       
   170 
       
   171 fun try_conv' cv ct = cv ct handle Fail => raise Unchanged;
       
   172 
       
   173 fun (cv1 then_conv' cv2) ct =
       
   174   let val eq = cv1 ct
       
   175   in Thm.transitive eq (cv2 (Thm.rhs_of eq)) handle Unchanged => eq
       
   176   end handle Unchanged => cv2 ct;
       
   177 
       
   178 fun first_conv' cvs = fold_rev (curry op else_conv') cvs no_conv';
       
   179 
       
   180 fun comb_conv' cv ct =
       
   181   let val (ct1, ct2) = Thm.dest_comb ct
       
   182   in
       
   183     let val eq1 = cv ct1
       
   184     in Thm.combination eq1 (cv ct2) handle Unchanged =>
       
   185        Thm.combination eq1 (Thm.reflexive ct2)
       
   186     end handle Unchanged =>
       
   187     let val eq2 = cv ct2
       
   188     in Thm.combination (Thm.reflexive ct1) eq2 end
       
   189   end handle CTERM _ => raise Fail;
       
   190 
       
   191 fun abs_conv' cv ctxt ct =
       
   192   (case Thm.term_of ct of
       
   193     Abs (x, _, _) =>
       
   194       let
       
   195         val ([u], ctxt') = Variable.variant_fixes ["u"] ctxt;
       
   196         val (v, ct') = Thm.dest_abs (SOME u) ct;
       
   197         val eq = cv (v, ctxt') ct';
       
   198       in Thm.abstract_rule x v eq end
       
   199   | _ => raise Fail);
       
   200 
       
   201 fun subc' conv ctxt = 
       
   202   comb_conv' (conv ctxt) else_conv'
       
   203   abs_conv' (conv o snd) ctxt else_conv'
       
   204   all_conv';
       
   205 
       
   206 fun botc' conv ctxt ct =
       
   207   (subc' (botc' conv) ctxt then_conv'
       
   208    try_conv' (conv then_conv' botc' conv ctxt)) ct
       
   209 
       
   210 fun rewr_conv' rule ct =
       
   211   Conv.rewr_conv rule ct handle CTERM _ => raise Fail;
       
   212 *}
       
   213 
       
   214 text {* Fully rewrite the term (32 secs on my laptop) *}
       
   215 
       
   216 ML {*
       
   217 timeit (fn () =>
       
   218   botc' (first_conv' (map rewr_conv' eqns)) @{context} ct)
       
   219 *}
       
   220 
       
   221 ML {*
       
   222 botc' (first_conv' (map rewr_conv' eqns)) @{context}
       
   223   @{cterm "\<forall>x y z. P (rev [x, y, z])"}
       
   224 *}
       
   225 
       
   226 
       
   227 
       
   228 section {* Bottom-up writing using skeletons *}
       
   229 
       
   230 text {*
       
   231   Motivation: avoid re-inspecting terms that are already
       
   232   in normal form
       
   233   Skeleton = rhs of last rewrite rule applied
       
   234   decomposed in parallel with the term to be normalized
       
   235   if skeleton is a variable, then the corresponding
       
   236   term must already be in normal form.
       
   237 *}
       
   238 
       
   239 ML {*
       
   240 infix 1 then_conv'';
       
   241 
       
   242 fun (cv1 then_conv'' cv2) cts =
       
   243   let val (eq1, skel1) = cv1 cts
       
   244   in 
       
   245     let val (eq2, skel2) = cv2 (Thm.rhs_of eq1, skel1)
       
   246     in (Thm.transitive eq1 eq2, skel2) end handle Unchanged => (eq1, skel1)
       
   247   end handle Unchanged => cv2 cts;
       
   248 
       
   249 val dummy_skel = Bound 0;
       
   250 
       
   251 fun comb_conv'' cv (ct, skel) =
       
   252   let
       
   253     val (ct1, ct2) = Thm.dest_comb ct
       
   254     val (skel1, skel2) = (case skel of
       
   255         skel1 $ skel2 => (skel1, skel2)
       
   256       | _ => (dummy_skel, dummy_skel));
       
   257   in
       
   258     let val (eq1, skel1') = cv (ct1, skel1)
       
   259     in 
       
   260       let val (eq2, skel2') = cv (ct2, skel2)
       
   261       in (Thm.combination eq1 eq2, skel1' $ skel2') end
       
   262       handle Unchanged =>
       
   263       (Thm.combination eq1 (Thm.reflexive ct2), skel1' $ skel2)
       
   264     end handle Unchanged =>
       
   265     let val (eq2, skel2') = cv (ct2, skel2)
       
   266     in (Thm.combination (Thm.reflexive ct1) eq2, skel1 $ skel2') end
       
   267   end handle CTERM _ => raise Fail;
       
   268 
       
   269 fun abs_conv'' cv ctxt (ct, skel) =
       
   270   (case Thm.term_of ct of
       
   271     Abs (x, T, _) =>
       
   272       let
       
   273         val ([u], ctxt') = Variable.variant_fixes ["u"] ctxt;
       
   274         val (v, ct') = Thm.dest_abs (SOME u) ct;
       
   275         val skel' = (case skel of
       
   276             Abs (_, _, skel') => skel'
       
   277           | _ => dummy_skel)
       
   278         val (eq, skel'') = cv (v, ctxt') (ct', skel');
       
   279       in (Thm.abstract_rule x v eq, Abs (x, T, skel'')) end
       
   280   | _ => raise Fail);
       
   281 
       
   282 fun subc'' conv ctxt =
       
   283   comb_conv'' (conv ctxt) else_conv'
       
   284   abs_conv'' (conv o snd) ctxt else_conv'
       
   285   all_conv';
       
   286 
       
   287 fun botc'' conv ctxt (_, Var _) = raise Unchanged
       
   288   | botc'' conv ctxt cts =
       
   289       (subc'' (botc'' conv) ctxt then_conv''
       
   290        try_conv' (conv then_conv'' botc'' conv ctxt)) cts
       
   291 
       
   292 fun rewr_conv'' rule (ct, _) =
       
   293   (Conv.rewr_conv rule ct, term_of (Thm.rhs_of rule))
       
   294   handle CTERM _ => raise Fail;
       
   295 *}
       
   296 
       
   297 text {* Fully rewrite the term (1.5 secs on my laptop) *}
       
   298 
       
   299 ML {*
       
   300 timeit (fn () =>
       
   301   fst (botc'' (first_conv' (map rewr_conv'' eqns)) @{context} (ct, dummy_skel)))
       
   302 *}
       
   303 
       
   304 ML {*
       
   305 fst (botc'' (first_conv' (map rewr_conv'' eqns)) @{context}
       
   306   (@{cterm "\<forall>x y z. P (rev [x, y, z])"}, dummy_skel))
       
   307 *}
       
   308 
       
   309 
       
   310 
       
   311 section {* The simplifier *}
       
   312 
       
   313 text {*
       
   314   The simplifier is a conversion itself, using many of
       
   315   the techniques just described.
       
   316 *}
       
   317 
       
   318 ML {*
       
   319 Simplifier.rewrite (HOL_basic_ss addsimps eqns) ct
       
   320 *}
       
   321 
       
   322 ML {*
       
   323 Simplifier.rewrite @{simpset}
       
   324   @{cterm "\<And>(x::int) (y::int). P x \<Longrightarrow> x = 42 \<Longrightarrow> Q x \<Longrightarrow> R (y + 0) \<Longrightarrow> S (1 * x)"}
       
   325 *}
       
   326 
       
   327 ML {*
       
   328 Simplifier.asm_rewrite @{simpset}
       
   329   @{cterm "\<And>(x::int) (y::int). P x \<Longrightarrow> x = 42 \<Longrightarrow> Q x \<Longrightarrow> R (y + 0) \<Longrightarrow> S (1 * x)"}
       
   330 *}
       
   331 
       
   332 ML {*
       
   333 Simplifier.full_rewrite @{simpset}
       
   334   @{cterm "\<And>(x::int) (y::int). P x \<Longrightarrow> x = 42 \<Longrightarrow> Q x \<Longrightarrow> R (y + 0) \<Longrightarrow> S (1 * x)"}
       
   335 *}
       
   336 
       
   337 ML {*
       
   338 Simplifier.asm_full_rewrite @{simpset}
       
   339   @{cterm "\<And>(x::int) (y::int). P x \<Longrightarrow> x = 42 \<Longrightarrow> Q x \<Longrightarrow> R (y + 0) \<Longrightarrow> S (1 * x)"}
       
   340 *}
       
   341 
       
   342 ML {*
       
   343 Simplifier.asm_lr_rewrite @{simpset}
       
   344   @{cterm "\<And>(x::int) (y::int). P x \<Longrightarrow> x = 42 \<Longrightarrow> Q x \<Longrightarrow> R (y + 0) \<Longrightarrow> S (1 * x)"}
       
   345 *}
       
   346 
       
   347 section {* Simplification procedures *}
       
   348 
       
   349 text {*
       
   350   A simplification procedure is a function of type
       
   351   @{ML_type "cterm -> thm option"}
       
   352   It can be used to prove rewrite rules on-the-fly.
       
   353 *}
       
   354 
       
   355 text {* Example 1: One-point rules *}
       
   356 
       
   357 text {*
       
   358   Problem: how to rewrite
       
   359   @{term "\<And>x y z. P y \<Longrightarrow> y = t \<Longrightarrow> Q y"}
       
   360   to
       
   361   @{term "\<And>x z. P t \<Longrightarrow> Q t"}
       
   362 *}
       
   363 
       
   364 lemma meta_onepoint1: "(\<And>x. x = t \<Longrightarrow> PROP P x) \<equiv> PROP P t"
       
   365 proof
       
   366   assume R: "\<And>x. x = t \<Longrightarrow> PROP P x"
       
   367   show "PROP P t" by (rule R [OF refl])
       
   368 next
       
   369   fix x assume "PROP P t" "x = t"
       
   370   then show "PROP P x" by simp
       
   371 qed
       
   372 
       
   373 lemma meta_onepoint2: "(\<And>x. t = x \<Longrightarrow> PROP P x) \<equiv> PROP P t"
       
   374 proof
       
   375   assume R: "\<And>x. t = x \<Longrightarrow> PROP P x"
       
   376   show "PROP P t" by (rule R [OF refl])
       
   377 next
       
   378   fix x assume "PROP P t" "t = x"
       
   379   then show "PROP P x" by simp
       
   380 qed
       
   381 
       
   382 lemmas meta_onepoint = meta_onepoint1 meta_onepoint2
       
   383 
       
   384 text {*
       
   385   Note: only works with formulae like
       
   386   @{term "\<And>x z y. y = t \<Longrightarrow> P y \<Longrightarrow> Q y"}
       
   387   but not
       
   388   @{term "\<And>x y z. P y \<Longrightarrow> y = t \<Longrightarrow> Q y"}
       
   389   Solution: reorder quantifiers and premises
       
   390 *}
       
   391 
       
   392 ML {*
       
   393 Simplifier.rewrite
       
   394   (HOL_basic_ss addsimps @{thms meta_onepoint})
       
   395   @{cterm "\<And>x z y. y = t \<Longrightarrow> P y \<Longrightarrow> Q y"}
       
   396 *}
       
   397 
       
   398 ML {*
       
   399 Simplifier.rewrite
       
   400   (HOL_basic_ss addsimps @{thms meta_onepoint})
       
   401   @{cterm "\<And>x y z. P y \<Longrightarrow> y = t \<Longrightarrow> Q y"}
       
   402 *}
       
   403 
       
   404 text {* Move parameters to the right *}
       
   405 
       
   406 ML {*
       
   407 fun swap_params_conv ctxt i j cv =
       
   408   let
       
   409     fun conv1 0 ctxt = Conv.forall_conv (cv o snd) ctxt
       
   410       | conv1 k ctxt =
       
   411           Conv.rewr_conv @{thm swap_params} then_conv
       
   412           Conv.forall_conv (conv1 (k-1) o snd) ctxt
       
   413     fun conv2 0 ctxt = conv1 j ctxt
       
   414       | conv2 k ctxt = Conv.forall_conv (conv2 (k-1) o snd) ctxt
       
   415   in conv2 i ctxt end;
       
   416 *}
       
   417 
       
   418 ML {*
       
   419 swap_params_conv @{context} 2 3 (K Conv.all_conv)
       
   420   @{cterm "\<And>a b c d e f. P a b c d e f"}
       
   421 *}
       
   422 
       
   423 text {* Move premises to the left *}
       
   424 
       
   425 ML {*
       
   426 fun swap_prems_conv 0 = Conv.all_conv
       
   427   | swap_prems_conv i =
       
   428       Conv.implies_concl_conv (swap_prems_conv (i-1)) then_conv
       
   429       Conv.rewr_conv Drule.swap_prems_eq
       
   430 *}
       
   431 
       
   432 ML {*
       
   433 swap_prems_conv 3
       
   434   @{cterm "A \<Longrightarrow> B \<Longrightarrow> C \<Longrightarrow> D \<Longrightarrow> E \<Longrightarrow> F"}
       
   435 *}
       
   436 
       
   437 text {* Find out which equation to move *}
       
   438 
       
   439 ML {*
       
   440 fun find_eq t =
       
   441   let
       
   442     val l = length (Logic.strip_params t);
       
   443     val Hs = Logic.strip_assums_hyp t;
       
   444     fun find (i, (_ $ (Const ("op =", _) $ Bound j $ _))) = SOME (i, j)
       
   445       | find (i, (_ $ (Const ("op =", _) $ _ $ Bound j))) = SOME (i, j)
       
   446       | find _ = NONE
       
   447   in
       
   448     case get_first find (map_index I Hs) of
       
   449       NONE => NONE
       
   450     | SOME (0, 0) => NONE
       
   451     | SOME (i, j) => SOME (i, l - j - 1, j)
       
   452   end;
       
   453 *}
       
   454 
       
   455 ML {*
       
   456 find_eq @{term "\<And>x y z. P y \<Longrightarrow> y = t \<Longrightarrow> Q y"}
       
   457 *}
       
   458 
       
   459 text {* Turn it into a simproc *}
       
   460 
       
   461 ML {*
       
   462 fun mk_rrule ctxt ct = case find_eq (term_of ct) of
       
   463     NONE => NONE
       
   464   | SOME (i, k, j) => SOME (swap_params_conv ctxt k j (K (swap_prems_conv i)) ct);
       
   465 *}
       
   466 
       
   467 ML {*
       
   468 val rearrange_eqs_simproc =
       
   469   Simplifier.simproc @{theory} "rearrange_eqs" ["all t"] (fn thy => fn ss => fn t =>
       
   470     mk_rrule (Simplifier.the_context ss) (cterm_of thy t))
       
   471 *}
       
   472 
       
   473 ML {*
       
   474 Simplifier.rewrite
       
   475   (HOL_basic_ss addsimps @{thms meta_onepoint}
       
   476      addsimprocs [rearrange_eqs_simproc])
       
   477   @{cterm "\<And>x y z. P y \<Longrightarrow> y = t \<Longrightarrow> Q y"}
       
   478 *}
       
   479 
       
   480 subsection {* Example 2: Simplifying set comprehensions *}
       
   481 
       
   482 text {*
       
   483   Problem: How to simplify
       
   484   @{term "{(x, y, z). (x, y, z) \<in> S}"}
       
   485   to
       
   486   @{term S}
       
   487 *}
       
   488 
       
   489 ML {* @{thm Collect_mem_eq} *}
       
   490 
       
   491 text {*
       
   492   Note: does not work with pairs
       
   493 *}
       
   494 
       
   495 ML {*
       
   496 Simplifier.rewrite
       
   497   (HOL_basic_ss addsimps @{thms Collect_mem_eq})
       
   498   @{cterm "P {x. x \<in> S}"}
       
   499 *}
       
   500 
       
   501 ML {*
       
   502 Simplifier.rewrite
       
   503   (HOL_basic_ss addsimps @{thms Collect_mem_eq})
       
   504   @{cterm "P {(x, y). (x, y) \<in> S}"}
       
   505 *}
       
   506 
       
   507 text {*
       
   508   Write a simproc to prove
       
   509   @{term "{(x, y). (x, y) \<in> S} \<equiv> S"}
       
   510 *}
       
   511 
       
   512 lemma test: "{(x, y). (x, y) \<in> S} \<equiv> S"
       
   513   apply (rule eq_reflection)
       
   514   apply (rule subset_antisym)
       
   515   apply (rule subsetI)
       
   516   apply (drule CollectD)
       
   517   apply (simp only: split_paired_all split_conv)
       
   518   apply (rule subsetI)
       
   519   apply (rule CollectI)
       
   520   apply (simp only: split_paired_all split_conv)
       
   521   done
       
   522 
       
   523 text {* The same in ML *}
       
   524 
       
   525 ML {*
       
   526 let
       
   527   val simp = full_simp_tac
       
   528     (HOL_basic_ss addsimps [split_paired_all, split_conv]) 1
       
   529 in
       
   530   Goal.prove @{context} [] []
       
   531     @{term "{(x, y). (x, y) \<in> S} \<equiv> S"}
       
   532     (K (EVERY
       
   533       [rtac eq_reflection 1, rtac @{thm subset_antisym} 1,
       
   534        rtac subsetI 1, dtac CollectD 1, simp,
       
   535        rtac subsetI 1, rtac CollectI 1, simp]))
       
   536 end
       
   537 *}
       
   538 
       
   539 ML {*
       
   540 val (u, Ts, ps) = HOLogic.strip_split
       
   541   @{term "\<lambda>(x, y). (x, y) \<in> S"}
       
   542 *}
       
   543 
       
   544 ML {*
       
   545 val collect_mem_simproc =
       
   546   Simplifier.simproc (theory "Set") "Collect_mem" ["Collect t"] (fn thy => fn ss =>
       
   547     fn S as Const ("Collect", Type ("fun", [_, T])) $ t =>
       
   548          let val (u, Ts, ps) = HOLogic.strip_split t
       
   549          in case u of
       
   550            (c as Const ("op :", _)) $ q $ S' =>
       
   551              (case try (HOLogic.dest_tuple' ps) q of
       
   552                 NONE => NONE
       
   553               | SOME ts =>
       
   554                   if not (loose_bvar (S', 0)) andalso
       
   555                     ts = map Bound (length ps downto 0)
       
   556                   then
       
   557                     let val simp = full_simp_tac (Simplifier.inherit_context ss
       
   558                       (HOL_basic_ss addsimps [split_paired_all, split_conv])) 1
       
   559                     in
       
   560                       SOME (Goal.prove (Simplifier.the_context ss) [] []
       
   561                         (Const ("==", T --> T --> propT) $ S $ S')
       
   562                         (K (EVERY
       
   563                           [rtac eq_reflection 1, rtac @{thm subset_antisym} 1,
       
   564                            rtac subsetI 1, dtac CollectD 1, simp,
       
   565                            rtac subsetI 1, rtac CollectI 1, simp])))
       
   566                     end
       
   567                   else NONE)
       
   568          | _ => NONE
       
   569          end
       
   570      | _ => NONE);
       
   571 *}
       
   572 
       
   573 ML {*
       
   574 Simplifier.rewrite
       
   575   (HOL_basic_ss addsimps @{thms Collect_mem_eq}
       
   576      addsimprocs [collect_mem_simproc])
       
   577   @{cterm "P {(x, y). (x, y) \<in> S}"}
       
   578 *}
       
   579 
       
   580 end