src/HOL/Tools/set_comprehension_pointfree.ML
author wenzelm
Sat, 30 Nov 2024 13:40:57 +0100
changeset 81511 8cbc8bc6f382
parent 80655 be3325cbeb40
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48124
87c831e30f0a ignore morphism more explicitly;
wenzelm
parents: 48109
diff changeset
     1
(*  Title:      HOL/Tools/set_comprehension_pointfree.ML
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     2
    Author:     Felix Kuperjans, Lukas Bulwahn, TU Muenchen
48124
87c831e30f0a ignore morphism more explicitly;
wenzelm
parents: 48109
diff changeset
     3
    Author:     Rafal Kolanski, NICTA
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     4
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     5
Simproc for rewriting set comprehensions to pointfree expressions.
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     6
*)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     7
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     8
signature SET_COMPREHENSION_POINTFREE =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     9
sig
78801
42ae6e0ecfd4 tuned signature;
wenzelm
parents: 78800
diff changeset
    10
  val base_proc : Simplifier.proc
42ae6e0ecfd4 tuned signature;
wenzelm
parents: 78800
diff changeset
    11
  val code_proc : Simplifier.proc
42ae6e0ecfd4 tuned signature;
wenzelm
parents: 78800
diff changeset
    12
  val proc : Simplifier.proc
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    13
end
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    14
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    15
structure Set_Comprehension_Pointfree : SET_COMPREHENSION_POINTFREE =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    16
struct
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    17
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    18
(* syntactic operations *)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    19
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    20
fun mk_inf (t1, t2) =
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    21
  let val T = fastype_of t1
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    22
  in \<^Const>\<open>inf T for t1 t2\<close> end
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    23
49768
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    24
fun mk_sup (t1, t2) =
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    25
  let val T = fastype_of t1
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    26
  in \<^Const>\<open>sup T for t1 t2\<close> end
49768
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    27
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    28
fun mk_Compl t =
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    29
  let val T = fastype_of t
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    30
  in \<^Const>\<open>uminus T for t\<close> end
49768
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    31
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    32
fun mk_image t1 t2 =
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    33
  let val \<^Type>\<open>fun A B\<close> = fastype_of t1
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    34
  in \<^Const>\<open>image A B for t1 t2\<close> end;
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    35
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    36
fun mk_sigma (t1, t2) =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    37
  let
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    38
    val \<^Type>\<open>set A\<close> = fastype_of t1
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    39
    val \<^Type>\<open>set B\<close> = fastype_of t2
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    40
  in \<^Const>\<open>Sigma A B for t1 \<open>absdummy A t2\<close>\<close> end;
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    41
49874
72b6d5fb407f term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
bulwahn
parents: 49873
diff changeset
    42
fun mk_vimage f s =
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    43
  let val \<^Type>\<open>fun A B\<close> = fastype_of f
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    44
  in \<^Const>\<open>vimage A B for f s\<close> end;
49874
72b6d5fb407f term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
bulwahn
parents: 49873
diff changeset
    45
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    46
fun dest_Collect \<^Const_>\<open>Collect _ for \<open>Abs (x, T, t)\<close>\<close> = ((x, T), t)
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    47
  | dest_Collect t = raise TERM ("dest_Collect", [t])
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    48
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    49
(* Copied from predicate_compile_aux.ML *)
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    50
fun strip_ex \<^Const_>\<open>Ex _ for \<open>Abs (x, T, t)\<close>\<close> =
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    51
      let val (xTs, t') = strip_ex t
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    52
      in ((x, T) :: xTs, t') end
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    53
  | strip_ex t = ([], t)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    54
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    55
fun mk_prod1 Ts (t1, t2) =
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    56
  let val (A, B) = apply2 (curry fastype_of1 Ts) (t1, t2)
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    57
  in \<^Const>\<open>Pair A B for t1 t2\<close> end;
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    58
49874
72b6d5fb407f term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
bulwahn
parents: 49873
diff changeset
    59
fun mk_split_abs vs (Bound i) t = let val (x, T) = nth vs i in Abs (x, T, t) end
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    60
  | mk_split_abs vs \<^Const_>\<open>Pair _ _ for u v\<close> t =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
    61
      HOLogic.mk_case_prod (mk_split_abs vs u (mk_split_abs vs v t))
49874
72b6d5fb407f term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
bulwahn
parents: 49873
diff changeset
    62
  | mk_split_abs _ t _ = raise TERM ("mk_split_abs: bad term", [t]);
72b6d5fb407f term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
bulwahn
parents: 49873
diff changeset
    63
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
    64
(* a variant of HOLogic.strip_ptupleabs *)
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
    65
val strip_ptupleabs =
49901
58cac1b3b535 comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
bulwahn
parents: 49900
diff changeset
    66
  let
58cac1b3b535 comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
bulwahn
parents: 49900
diff changeset
    67
    fun strip [] qs vs t = (t, rev vs, qs)
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    68
      | strip (p :: ps) qs vs \<^Const_>\<open>case_prod _ _ _ for t\<close> =
49901
58cac1b3b535 comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
bulwahn
parents: 49900
diff changeset
    69
          strip ((1 :: p) :: (2 :: p) :: ps) (p :: qs) vs t
49957
6250121bfffb passing around the simpset instead of the context; rewriting tactics to avoid the 'renamed bound variable' warnings in nested simplifier calls
bulwahn
parents: 49946
diff changeset
    70
      | strip (_ :: ps) qs vs (Abs (s, T, t)) = strip ps qs ((s, T) :: vs) t
6250121bfffb passing around the simpset instead of the context; rewriting tactics to avoid the 'renamed bound variable' warnings in nested simplifier calls
bulwahn
parents: 49946
diff changeset
    71
      | strip (_ :: ps) qs vs t = strip ps qs
49901
58cac1b3b535 comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
bulwahn
parents: 49900
diff changeset
    72
          ((Name.uu_, hd (binder_types (fastype_of1 (map snd vs, t)))) :: vs)
58cac1b3b535 comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
bulwahn
parents: 49900
diff changeset
    73
          (incr_boundvars 1 t $ Bound 0)
58cac1b3b535 comprehension conversion reuses suggested names for bound variables instead of invented fresh ones; tuned tactic
bulwahn
parents: 49900
diff changeset
    74
  in strip [[]] [] [] end;
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    75
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    76
(* patterns *)
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    77
50024
b7265db3a1dc simplified structure of patterns in set_comprehension_simproc
bulwahn
parents: 49959
diff changeset
    78
datatype pattern = Pattern of int list
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    79
50024
b7265db3a1dc simplified structure of patterns in set_comprehension_simproc
bulwahn
parents: 49959
diff changeset
    80
fun dest_Pattern (Pattern bs) = bs
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    81
50024
b7265db3a1dc simplified structure of patterns in set_comprehension_simproc
bulwahn
parents: 49959
diff changeset
    82
fun type_of_pattern Ts (Pattern bs) = HOLogic.mk_tupleT (map (nth Ts) bs)
b7265db3a1dc simplified structure of patterns in set_comprehension_simproc
bulwahn
parents: 49959
diff changeset
    83
b7265db3a1dc simplified structure of patterns in set_comprehension_simproc
bulwahn
parents: 49959
diff changeset
    84
fun term_of_pattern Ts (Pattern bs) =
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    85
    let
50024
b7265db3a1dc simplified structure of patterns in set_comprehension_simproc
bulwahn
parents: 49959
diff changeset
    86
      fun mk [b] = Bound b
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    87
        | mk (b :: bs) =
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    88
            \<^Const>\<open>Pair \<open>nth Ts b\<close> \<open>type_of_pattern Ts (Pattern bs)\<close> for \<open>Bound b\<close> \<open>mk bs\<close>\<close>
50024
b7265db3a1dc simplified structure of patterns in set_comprehension_simproc
bulwahn
parents: 49959
diff changeset
    89
    in mk bs end;
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    90
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    91
(* formulas *)
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    92
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    93
datatype formula = Atom of (pattern * term) | Int of formula * formula | Un of formula * formula
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
    94
49900
89b118c0c070 checking for bound variables in the set expression; handling negation more generally
bulwahn
parents: 49898
diff changeset
    95
fun map_atom f (Atom a) = Atom (f a)
89b118c0c070 checking for bound variables in the set expression; handling negation more generally
bulwahn
parents: 49898
diff changeset
    96
  | map_atom _ x = x
89b118c0c070 checking for bound variables in the set expression; handling negation more generally
bulwahn
parents: 49898
diff changeset
    97
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    98
fun is_collect_atom (Atom (_, \<^Const_>\<open>Collect _ for _\<close>)) = true
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
    99
  | is_collect_atom (Atom (_, \<^Const_>\<open>uminus _ for \<^Const_>\<open>Collect _ for _\<close>\<close>)) = true
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   100
  | is_collect_atom _ = false
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   101
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
   102
fun mk_case_prod _ [(x, T)] t = (T, Abs (x, T, t))
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
   103
  | mk_case_prod rT ((x, T) :: vs) t =
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   104
    let
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
   105
      val (T', t') = mk_case_prod rT vs t
80642
wenzelm
parents: 80641
diff changeset
   106
      val t'' = \<^Const>\<open>case_prod T T' rT for \<open>Abs (x, T, t')\<close>\<close>
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   107
    in (domain_type (fastype_of t''), t'') end
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   108
50030
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   109
fun mk_term vs t =
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   110
  let
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   111
    val bs = loose_bnos t
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   112
    val vs' = map (nth (rev vs)) bs
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   113
    val subst = map_index (fn (i, j) => (j, Bound i)) (rev bs)
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   114
      |> sort (fn (p1, p2) => int_ord (fst p1, fst p2))
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   115
      |> (fn subst' => map (fn i => the_default (Bound i) (AList.lookup (op =) subst' i)) (0 upto (fst (snd (split_last subst')))))
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   116
    val t' = subst_bounds (subst, t)
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   117
    val tuple = Pattern bs
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   118
  in (tuple, (vs', t')) end
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   119
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   120
fun default_atom vs t =
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   121
  let
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   122
    val (tuple, (vs', t')) = mk_term vs t
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   123
    val T = HOLogic.mk_tupleT (map snd vs')
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   124
    val s = \<^Const>\<open>Collect T for \<open>snd (mk_case_prod \<^Type>\<open>bool\<close> vs' t')\<close>\<close>
50030
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   125
  in
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   126
    (tuple, Atom (tuple, s))
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   127
  end
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   128
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   129
fun mk_atom vs (t as \<^Const_>\<open>Set.member _ for x s\<close>) =
49900
89b118c0c070 checking for bound variables in the set expression; handling negation more generally
bulwahn
parents: 49898
diff changeset
   130
    if not (null (loose_bnos s)) then
50032
a439a9d14ba3 handling x : S y pattern with the default mechanism instead of raising an exception in the set_comprehension_pointfree simproc
bulwahn
parents: 50031
diff changeset
   131
      default_atom vs t
49900
89b118c0c070 checking for bound variables in the set expression; handling negation more generally
bulwahn
parents: 49898
diff changeset
   132
    else
80642
wenzelm
parents: 80641
diff changeset
   133
      (case try (map (fn Bound i => i) o HOLogic.strip_tuple) x of
50030
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   134
      SOME pat => (Pattern pat, Atom (Pattern pat, s))
49874
72b6d5fb407f term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
bulwahn
parents: 49873
diff changeset
   135
    | NONE =>
49900
89b118c0c070 checking for bound variables in the set expression; handling negation more generally
bulwahn
parents: 49898
diff changeset
   136
        let
80640
wenzelm
parents: 78801
diff changeset
   137
          val (tuple, (vs', x')) = mk_term vs x
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   138
          val \<^Type>\<open>set rT\<close> = fastype_of s
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
   139
          val s = mk_vimage (snd (mk_case_prod rT vs' x')) s
50030
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   140
        in (tuple, Atom (tuple, s)) end)
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   141
  | mk_atom vs \<^Const_>\<open>Not for t\<close> = apsnd (map_atom (apsnd mk_Compl)) (mk_atom vs t)
50030
349f651ec203 syntactic tuning and restructuring of set_comprehension_pointfree simproc
bulwahn
parents: 50029
diff changeset
   142
  | mk_atom vs t = default_atom vs t
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   143
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   144
fun merge' [] (pats1, pats2) = ([], (pats1, pats2))
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   145
  | merge' pat (pats, []) = (pat, (pats, []))
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   146
  | merge' pat (pats1, pats) =
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   147
  let
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   148
    fun disjoint_to_pat p = null (inter (op =) pat p)
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   149
    val overlap_pats = filter_out disjoint_to_pat pats
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   150
    val rem_pats = filter disjoint_to_pat pats
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   151
    val (pat, (pats', pats1')) = merge' (distinct (op =) (flat overlap_pats @ pat)) (rem_pats, pats1)
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   152
  in
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   153
    (pat, (pats1', pats'))
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   154
  end
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   155
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   156
fun merge ([], pats) = pats
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   157
  | merge (pat :: pats', pats) =
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   158
  let val (pat', (pats1', pats2')) = merge' pat (pats', pats)
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   159
  in pat' :: merge (pats1', pats2') end;
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   160
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   161
fun restricted_merge ([], pats) = pats
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   162
  | restricted_merge (pat :: pats', pats) =
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   163
  let
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   164
    fun disjoint_to_pat p = null (inter (op =) pat p)
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   165
    val overlap_pats = filter_out disjoint_to_pat pats
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   166
    val rem_pats = filter disjoint_to_pat pats
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   167
  in
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   168
    case overlap_pats of
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   169
      [] => pat :: restricted_merge (pats', rem_pats)
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   170
    | [pat'] => if subset (op =) (pat, pat') then
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   171
        pat' :: restricted_merge (pats', rem_pats)
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   172
      else if subset (op =) (pat', pat) then
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   173
        pat :: restricted_merge (pats', rem_pats)
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   174
      else error "restricted merge: two patterns require relational join"
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   175
    | _ => error "restricted merge: multiple patterns overlap"
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   176
  end;
80640
wenzelm
parents: 78801
diff changeset
   177
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   178
fun map_atoms f (Atom a) = Atom (f a)
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58963
diff changeset
   179
  | map_atoms f (Un (fm1, fm2)) = Un (apply2 (map_atoms f) (fm1, fm2))
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58963
diff changeset
   180
  | map_atoms f (Int (fm1, fm2)) = Int (apply2 (map_atoms f) (fm1, fm2))
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   181
50028
d05f859558a0 improving the extension of sets in case of more than one bound variable; rearranging the tactic to prefer simpler steps before more involved ones
bulwahn
parents: 50026
diff changeset
   182
fun extend Ts bs t = foldr1 mk_sigma (t :: map (fn b => HOLogic.mk_UNIV (nth Ts b)) bs)
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   183
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   184
fun rearrange vs (pat, pat') t =
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   185
  let
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   186
    val subst = map_index (fn (i, b) => (b, i)) (rev pat)
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   187
    val vs' = map (nth (rev vs)) pat
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   188
    val Ts' = map snd (rev vs')
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   189
    val bs = map (fn b => the (AList.lookup (op =) subst b)) pat'
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   190
    val rt = term_of_pattern Ts' (Pattern bs)
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   191
    val rT = type_of_pattern Ts' (Pattern bs)
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
   192
    val (_, f) = mk_case_prod rT vs' rt
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   193
  in
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   194
    mk_image f t
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   195
  end;
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   196
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   197
fun adjust vs pats (Pattern pat, t) =
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   198
  let
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   199
    val SOME p = find_first (fn p => not (null (inter (op =) pat p))) pats
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   200
    val missing = subtract (op =) pat p
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   201
    val Ts = rev (map snd vs)
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   202
    val t' = extend Ts missing t
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   203
  in (Pattern p, rearrange vs (pat @ missing, p) t') end
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   204
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   205
fun adjust_atoms vs pats fm = map_atoms (adjust vs pats) fm
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   206
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   207
fun merge_inter vs (pats1, fm1) (pats2, fm2) =
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   208
  let
80640
wenzelm
parents: 78801
diff changeset
   209
    val pats = restricted_merge (map dest_Pattern pats1, map dest_Pattern pats2)
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58963
diff changeset
   210
    val (fm1', fm2') = apply2 (adjust_atoms vs pats) (fm1, fm2)
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   211
  in
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   212
    (map Pattern pats, Int (fm1', fm2'))
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   213
  end;
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   214
80640
wenzelm
parents: 78801
diff changeset
   215
fun merge_union vs (pats1, fm1) (pats2, fm2) =
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   216
  let
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   217
    val pats = merge (map dest_Pattern pats1, map dest_Pattern pats2)
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58963
diff changeset
   218
    val (fm1', fm2') = apply2 (adjust_atoms vs pats) (fm1, fm2)
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   219
  in
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   220
    (map Pattern pats, Un (fm1', fm2'))
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   221
  end;
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   222
74383
107941e8fa01 clarified antiquotations;
wenzelm
parents: 70326
diff changeset
   223
fun mk_formula vs \<^Const_>\<open>conj for t1 t2\<close> = merge_inter vs (mk_formula vs t1) (mk_formula vs t2)
107941e8fa01 clarified antiquotations;
wenzelm
parents: 70326
diff changeset
   224
  | mk_formula vs \<^Const_>\<open>disj for t1 t2\<close> = merge_union vs (mk_formula vs t1) (mk_formula vs t2)
49874
72b6d5fb407f term construction of set_comprehension_pointfree simproc handles f x y : S patterns with Set.vimage
bulwahn
parents: 49873
diff changeset
   225
  | mk_formula vs t = apfst single (mk_atom vs t)
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   226
80640
wenzelm
parents: 78801
diff changeset
   227
fun strip_Int (Int (fm1, fm2)) = fm1 :: (strip_Int fm2)
49852
caaa1956f0da refined tactic in set_comprehension_pointfree simproc
bulwahn
parents: 49850
diff changeset
   228
  | strip_Int fm = [fm]
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   229
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   230
(* term construction *)
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   231
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   232
fun reorder_bounds pats t =
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   233
  let
50024
b7265db3a1dc simplified structure of patterns in set_comprehension_simproc
bulwahn
parents: 49959
diff changeset
   234
    val bounds = maps dest_Pattern pats
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   235
    val bperm = bounds ~~ ((length bounds - 1) downto 0)
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   236
      |> sort (fn (i,j) => int_ord (fst i, fst j)) |> map snd
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   237
  in
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   238
    subst_bounds (map Bound bperm, t)
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   239
  end;
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   240
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   241
fun is_reordering t =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
   242
  let val (t', _, _) = HOLogic.strip_ptupleabs t
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   243
  in forall (fn Bound _ => true) (HOLogic.strip_tuple t') end
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   244
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   245
fun mk_pointfree_expr t =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   246
  let
49857
7bf407d77152 setcomprehension_pointfree simproc also works for set comprehension without an equation
bulwahn
parents: 49852
diff changeset
   247
    val ((x, T), (vs, t'')) = apsnd strip_ex (dest_Collect t)
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   248
    val Ts = map snd (rev vs)
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   249
    fun mk_mem_UNIV n = HOLogic.mk_mem (Bound n, HOLogic.mk_UNIV (nth Ts n))
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   250
    fun lookup (pat', t) pat = if pat = pat' then t else HOLogic.mk_UNIV (type_of_pattern Ts pat)
49761
b7772f3b6c03 set_comprehension_pointfree also handles terms where the equation is not at the first position, which is a necessary generalisation to eventually handle bounded existentials; tuned
bulwahn
parents: 48128
diff changeset
   251
    val conjs = HOLogic.dest_conj t''
80642
wenzelm
parents: 80641
diff changeset
   252
    val refl = \<^Const>\<open>HOL.eq T for \<open>Bound (length vs)\<close> \<open>Bound (length vs)\<close>\<close>
49761
b7772f3b6c03 set_comprehension_pointfree also handles terms where the equation is not at the first position, which is a necessary generalisation to eventually handle bounded existentials; tuned
bulwahn
parents: 48128
diff changeset
   253
    val is_the_eq =
b7772f3b6c03 set_comprehension_pointfree also handles terms where the equation is not at the first position, which is a necessary generalisation to eventually handle bounded existentials; tuned
bulwahn
parents: 48128
diff changeset
   254
      the_default false o (try (fn eq => fst (HOLogic.dest_eq eq) = Bound (length vs)))
49857
7bf407d77152 setcomprehension_pointfree simproc also works for set comprehension without an equation
bulwahn
parents: 49852
diff changeset
   255
    val eq = the_default refl (find_first is_the_eq conjs)
49761
b7772f3b6c03 set_comprehension_pointfree also handles terms where the equation is not at the first position, which is a necessary generalisation to eventually handle bounded existentials; tuned
bulwahn
parents: 48128
diff changeset
   256
    val f = snd (HOLogic.dest_eq eq)
b7772f3b6c03 set_comprehension_pointfree also handles terms where the equation is not at the first position, which is a necessary generalisation to eventually handle bounded existentials; tuned
bulwahn
parents: 48128
diff changeset
   257
    val conjs' = filter_out (fn t => eq = t) conjs
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   258
    val unused_bounds = subtract (op =) (distinct (op =) (maps loose_bnos conjs'))
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   259
      (0 upto (length vs - 1))
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   260
    val (pats, fm) =
49943
6a26fed71755 passing names and types of all bounds around in the simproc
bulwahn
parents: 49942
diff changeset
   261
      mk_formula ((x, T) :: vs) (foldr1 HOLogic.mk_conj (conjs' @ map mk_mem_UNIV unused_bounds))
50031
bulwahn
parents: 50030
diff changeset
   262
    fun mk_set (Atom pt) = foldr1 mk_sigma (map (lookup pt) pats)
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   263
      | mk_set (Un (f1, f2)) = mk_sup (mk_set f1, mk_set f2)
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   264
      | mk_set (Int (f1, f2)) = mk_inf (mk_set f1, mk_set f2)
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   265
    val pat = foldr1 (mk_prod1 Ts) (map (term_of_pattern Ts) pats)
49857
7bf407d77152 setcomprehension_pointfree simproc also works for set comprehension without an equation
bulwahn
parents: 49852
diff changeset
   266
    val t = mk_split_abs (rev ((x, T) :: vs)) pat (reorder_bounds pats f)
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   267
  in
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   268
    if the_default false (try is_reordering t) andalso is_collect_atom fm then
80640
wenzelm
parents: 78801
diff changeset
   269
      error "mk_pointfree_expr: trivial case"
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   270
    else (fm, mk_image t (mk_set fm))
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   271
  end;
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   272
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   273
val rewrite_term = try mk_pointfree_expr
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   274
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   275
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   276
(* proof tactic *)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   277
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   278
val case_prod_beta = @{lemma "case_prod g x z = case_prod (\<lambda>x y. (g x y) z) x" by (simp add: case_prod_beta)}
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   279
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   280
val vimageI2' = @{lemma "f a \<notin> A ==> a \<notin> f -` A" by simp}
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   281
val vimageE' =
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   282
  @{lemma "a \<notin> f -` B ==> (\<And> x. f a = x ==> x \<notin> B ==> P) ==> P" by simp}
49944
28cd3c9ca278 tuned tactic in set_comprehension_pointfree simproc to handle composition of negation and vimage
bulwahn
parents: 49943
diff changeset
   283
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   284
val collectI' = @{lemma "\<not> P a ==> a \<notin> {x. P x}" by auto}
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   285
val collectE' = @{lemma "a \<notin> {x. P x} ==> (\<not> P a ==> Q) ==> Q" by auto}
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   286
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   287
fun elim_Collect_tac ctxt =
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   288
  dresolve_tac ctxt @{thms iffD1 [OF mem_Collect_eq]}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   289
  THEN' (REPEAT_DETERM o (eresolve_tac ctxt @{thms exE}))
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   290
  THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms conjE}
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
   291
  THEN' TRY o hyp_subst_tac ctxt;
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   292
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   293
fun intro_image_tac ctxt =
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   294
  resolve_tac ctxt @{thms image_eqI}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   295
  THEN' (REPEAT_DETERM1 o
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   296
      (resolve_tac ctxt @{thms refl}
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 61424
diff changeset
   297
      ORELSE' resolve_tac ctxt @{thms arg_cong2 [OF refl, where f = "(=)", OF prod.case, THEN iffD2]}
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   298
      ORELSE' CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1
51315
536a5603a138 provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents: 51314
diff changeset
   299
        (HOLogic.Trueprop_conv
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
   300
          (HOLogic.eq_conv Conv.all_conv (Conv.rewr_conv (mk_meta_eq case_prod_beta)))))) ctxt)))
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   301
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   302
fun elim_image_tac ctxt =
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   303
  eresolve_tac ctxt @{thms imageE}
50028
d05f859558a0 improving the extension of sets in case of more than one bound variable; rearranging the tactic to prefer simpler steps before more involved ones
bulwahn
parents: 50026
diff changeset
   304
  THEN' REPEAT_DETERM o CHANGED o
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55414
diff changeset
   305
    (TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_paired_all prod.case})
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   306
    THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject}
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
   307
    THEN' TRY o hyp_subst_tac ctxt)
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   308
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   309
fun tac1_of_formula ctxt (Int (fm1, fm2)) =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   310
    TRY o eresolve_tac ctxt @{thms conjE}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   311
    THEN' resolve_tac ctxt @{thms IntI}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   312
    THEN' (fn i => tac1_of_formula ctxt fm2 (i + 1))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   313
    THEN' tac1_of_formula ctxt fm1
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   314
  | tac1_of_formula ctxt (Un (fm1, fm2)) =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   315
    eresolve_tac ctxt @{thms disjE} THEN' resolve_tac ctxt @{thms UnI1}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   316
    THEN' tac1_of_formula ctxt fm1
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   317
    THEN' resolve_tac ctxt @{thms UnI2}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   318
    THEN' tac1_of_formula ctxt fm2
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   319
  | tac1_of_formula ctxt (Atom _) =
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   320
    REPEAT_DETERM1 o (assume_tac ctxt
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   321
      ORELSE' resolve_tac ctxt @{thms SigmaI}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   322
      ORELSE' ((resolve_tac ctxt @{thms CollectI} ORELSE' resolve_tac ctxt [collectI']) THEN'
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55414
diff changeset
   323
        TRY o simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}]))
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   324
      ORELSE' ((resolve_tac ctxt @{thms vimageI2} ORELSE' resolve_tac ctxt [vimageI2']) THEN'
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55414
diff changeset
   325
        TRY o simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}]))
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   326
      ORELSE' (resolve_tac ctxt @{thms image_eqI} THEN'
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   327
    (REPEAT_DETERM o
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   328
      (resolve_tac ctxt @{thms refl}
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 61424
diff changeset
   329
      ORELSE' resolve_tac ctxt @{thms arg_cong2[OF refl, where f = "(=)", OF prod.case, THEN iffD2]})))
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   330
      ORELSE' resolve_tac ctxt @{thms UNIV_I}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   331
      ORELSE' resolve_tac ctxt @{thms iffD2[OF Compl_iff]}
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   332
      ORELSE' assume_tac ctxt)
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   333
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   334
fun tac2_of_formula ctxt (Int (fm1, fm2)) =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   335
    TRY o eresolve_tac ctxt @{thms IntE}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   336
    THEN' TRY o resolve_tac ctxt @{thms conjI}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   337
    THEN' (fn i => tac2_of_formula ctxt fm2 (i + 1))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   338
    THEN' tac2_of_formula ctxt fm1
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   339
  | tac2_of_formula ctxt (Un (fm1, fm2)) =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   340
    eresolve_tac ctxt @{thms UnE} THEN' resolve_tac ctxt @{thms disjI1}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   341
    THEN' tac2_of_formula ctxt fm1
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   342
    THEN' resolve_tac ctxt @{thms disjI2}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   343
    THEN' tac2_of_formula ctxt fm2
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   344
  | tac2_of_formula ctxt (Atom _) =
50025
19965e6a705e handling arbitrary terms in the set comprehension and more general merging of patterns possible in the set_comprehension_pointfree simproc
bulwahn
parents: 50024
diff changeset
   345
    REPEAT_DETERM o
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   346
      (assume_tac ctxt
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   347
       ORELSE' dresolve_tac ctxt @{thms iffD1[OF mem_Sigma_iff]}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   348
       ORELSE' eresolve_tac ctxt @{thms conjE}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   349
       ORELSE' ((eresolve_tac ctxt @{thms CollectE} ORELSE' eresolve_tac ctxt [collectE']) THEN'
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55414
diff changeset
   350
         TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}]) THEN'
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   351
         REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject} THEN' TRY o hyp_subst_tac ctxt THEN'
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   352
         TRY o resolve_tac ctxt @{thms refl})
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   353
       ORELSE' (eresolve_tac ctxt @{thms imageE}
50028
d05f859558a0 improving the extension of sets in case of more than one bound variable; rearranging the tactic to prefer simpler steps before more involved ones
bulwahn
parents: 50026
diff changeset
   354
         THEN' (REPEAT_DETERM o CHANGED o
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55414
diff changeset
   355
         (TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_paired_all prod.case})
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   356
         THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   357
         THEN' TRY o hyp_subst_tac ctxt THEN' TRY o resolve_tac ctxt @{thms refl})))
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   358
       ORELSE' eresolve_tac ctxt @{thms ComplE}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   359
       ORELSE' ((eresolve_tac ctxt @{thms vimageE} ORELSE' eresolve_tac ctxt [vimageE'])
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55414
diff changeset
   360
        THEN' TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}])
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   361
        THEN' TRY o hyp_subst_tac ctxt THEN' TRY o resolve_tac ctxt @{thms refl}))
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   362
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   363
fun tac ctxt fm =
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   364
  let
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   365
    val subset_tac1 = resolve_tac ctxt @{thms subsetI}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   366
      THEN' elim_Collect_tac ctxt
50029
31c9294eebe6 using more proper simpset in tactic of set_comprehension_pointfree simproc to avoid renamed bound variable warnings in recursive simplifier calls
bulwahn
parents: 50028
diff changeset
   367
      THEN' intro_image_tac ctxt
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   368
      THEN' tac1_of_formula ctxt fm
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   369
    val subset_tac2 = resolve_tac ctxt @{thms subsetI}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   370
      THEN' elim_image_tac ctxt
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   371
      THEN' resolve_tac ctxt @{thms iffD2[OF mem_Collect_eq]}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   372
      THEN' REPEAT_DETERM o resolve_tac ctxt @{thms exI}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   373
      THEN' (TRY o REPEAT_ALL_NEW (resolve_tac ctxt @{thms conjI}))
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   374
      THEN' (K (TRY (FIRSTGOAL ((TRY o hyp_subst_tac ctxt) THEN' resolve_tac ctxt @{thms refl}))))
49852
caaa1956f0da refined tactic in set_comprehension_pointfree simproc
bulwahn
parents: 49850
diff changeset
   375
      THEN' (fn i => EVERY (rev (map_index (fn (j, f) =>
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   376
        REPEAT_DETERM (eresolve_tac ctxt @{thms IntE} (i + j)) THEN
58839
ccda99401bc8 eliminated aliases;
wenzelm
parents: 55642
diff changeset
   377
        tac2_of_formula ctxt f (i + j)) (strip_Int fm))))
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   378
  in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   379
    resolve_tac ctxt @{thms subset_antisym} THEN' subset_tac1 THEN' subset_tac2
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   380
  end;
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   381
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   382
49896
a39deedd5c3f employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
bulwahn
parents: 49875
diff changeset
   383
(* preprocessing conversion:
a39deedd5c3f employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
bulwahn
parents: 49875
diff changeset
   384
  rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} *)
a39deedd5c3f employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
bulwahn
parents: 49875
diff changeset
   385
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   386
fun comprehension_conv ctxt ct =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   387
  let
80642
wenzelm
parents: 80641
diff changeset
   388
    fun list_ex vs t = fold_rev (fn (x, T) => fn t => \<^Const>\<open>Ex T for \<open>Abs (x, T, t)\<close>\<close>) vs t
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   389
    fun mk_term t =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   390
      let
80642
wenzelm
parents: 80641
diff changeset
   391
        val \<^Const_>\<open>Collect T for t'\<close> = t
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61125
diff changeset
   392
        val (t'', vs, fp) = case strip_ptupleabs t' of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   393
            (_, [_], _) => raise TERM("mk_term", [t'])
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   394
          | (t'', vs, fp) => (t'', vs, fp)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   395
        val Ts = map snd vs
80642
wenzelm
parents: 80641
diff changeset
   396
        val eq =
wenzelm
parents: 80641
diff changeset
   397
          \<^Const>\<open>HOL.eq T for \<open>Bound (length Ts)\<close>
wenzelm
parents: 80641
diff changeset
   398
            \<open>HOLogic.mk_ptuple fp (HOLogic.mk_ptupleT fp Ts) (rev (map_index (Bound o #1) Ts))\<close>\<close>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   399
      in
80641
cc205e40192e tuned: more antiquotations;
wenzelm
parents: 80640
diff changeset
   400
        \<^Const>\<open>Collect T for \<open>absdummy T (list_ex vs (HOLogic.mk_conj (eq, t'')))\<close>\<close>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   401
      end;
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   402
    fun is_eq th = is_some (try (HOLogic.dest_eq o HOLogic.dest_Trueprop) (Thm.prop_of th))
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55414
diff changeset
   403
    val unfold_thms = @{thms split_paired_all mem_Collect_eq prod.case}
80640
wenzelm
parents: 78801
diff changeset
   404
    fun tac ctxt =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   405
      resolve_tac ctxt @{thms set_eqI}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   406
      THEN' simp_tac (put_simpset HOL_basic_ss ctxt addsimps unfold_thms)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   407
      THEN' resolve_tac ctxt @{thms iffI}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   408
      THEN' REPEAT_DETERM o resolve_tac ctxt @{thms exI}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   409
      THEN' resolve_tac ctxt @{thms conjI} THEN' resolve_tac ctxt @{thms refl} THEN' assume_tac ctxt
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   410
      THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms exE}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   411
      THEN' eresolve_tac ctxt @{thms conjE}
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   412
      THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject}
60696
8304fb4fb823 clarified context;
wenzelm
parents: 59621
diff changeset
   413
      THEN' Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
8304fb4fb823 clarified context;
wenzelm
parents: 59621
diff changeset
   414
        simp_tac (put_simpset HOL_basic_ss ctxt' addsimps (filter is_eq prems)) 1) ctxt
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   415
      THEN' TRY o assume_tac ctxt
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   416
  in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   417
    case try mk_term (Thm.term_of ct) of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   418
      NONE => Thm.reflexive ct
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   419
    | SOME t' =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   420
      Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (Thm.term_of ct, t')))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   421
          (fn {context, ...} => tac context 1)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   422
        RS @{thm eq_reflection}
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   423
  end
49896
a39deedd5c3f employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
bulwahn
parents: 49875
diff changeset
   424
a39deedd5c3f employing a preprocessing conversion that rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} in set_comprehension_pointfree simproc
bulwahn
parents: 49875
diff changeset
   425
49849
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   426
(* main simprocs *)
d9822ec4f434 extending the setcomprehension_pointfree simproc to handle nesting disjunctions, conjunctions and negations (with contributions from Rafal Kolanski, NICTA); tuned
bulwahn
parents: 49831
diff changeset
   427
49942
50e457bbc5fe locally inverting previously applied simplifications with ex_simps in set_comprehension_pointfree
bulwahn
parents: 49901
diff changeset
   428
val prep_thms =
50e457bbc5fe locally inverting previously applied simplifications with ex_simps in set_comprehension_pointfree
bulwahn
parents: 49901
diff changeset
   429
  map mk_meta_eq ([@{thm Bex_def}, @{thm Pow_iff[symmetric]}] @ @{thms ex_simps[symmetric]})
49873
4b7c2e4991fc extending preprocessing of simproc to rewrite subset inequality into membership of powerset
bulwahn
parents: 49857
diff changeset
   430
49850
873fa7156468 adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
bulwahn
parents: 49849
diff changeset
   431
val post_thms =
873fa7156468 adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
bulwahn
parents: 49849
diff changeset
   432
  map mk_meta_eq [@{thm Times_Un_distrib1[symmetric]},
873fa7156468 adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
bulwahn
parents: 49849
diff changeset
   433
  @{lemma "A \<times> B \<union> A \<times> C = A \<times> (B \<union> C)" by auto},
873fa7156468 adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
bulwahn
parents: 49849
diff changeset
   434
  @{lemma "(A \<times> B \<inter> C \<times> D) = (A \<inter> C) \<times> (B \<inter> D)" by auto}]
873fa7156468 adding postprocessing of computed pointfree expression in set_comprehension_pointfree simproc
bulwahn
parents: 49849
diff changeset
   435
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   436
fun conv ctxt t =
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   437
  let
70326
wenzelm
parents: 69597
diff changeset
   438
    val (t', ctxt') = yield_singleton (Variable.import_terms true) t (Variable.declare_term t ctxt)
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   439
    val ct = Thm.cterm_of ctxt' t'
49957
6250121bfffb passing around the simpset instead of the context; rewriting tactics to avoid the 'renamed bound variable' warnings in nested simplifier calls
bulwahn
parents: 49946
diff changeset
   440
    fun unfold_conv thms =
6250121bfffb passing around the simpset instead of the context; rewriting tactics to avoid the 'renamed bound variable' warnings in nested simplifier calls
bulwahn
parents: 49946
diff changeset
   441
      Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   442
        (empty_simpset ctxt' addsimps thms)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   443
    val prep_eq = (comprehension_conv ctxt' then_conv unfold_conv prep_thms) ct
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   444
    val t'' = Thm.term_of (Thm.rhs_of prep_eq)
50026
d9871e5ea0e1 importing term with schematic type variables properly before passing it to the tactic in the set_comprehension_pointfree simproc
bulwahn
parents: 50025
diff changeset
   445
    fun mk_thm (fm, t''') = Goal.prove ctxt' [] []
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   446
      (HOLogic.mk_Trueprop (HOLogic.mk_eq (t'', t'''))) (fn {context, ...} => tac context fm 1)
67710
cc2db3239932 added HOLogic.mk_obj_eq convenience and eliminated some clones;
wenzelm
parents: 67399
diff changeset
   447
    fun unfold th = th RS (HOLogic.mk_obj_eq prep_eq RS @{thm trans})
51315
536a5603a138 provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents: 51314
diff changeset
   448
    val post =
536a5603a138 provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents: 51314
diff changeset
   449
      Conv.fconv_rule
536a5603a138 provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents: 51314
diff changeset
   450
        (HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (unfold_conv post_thms)))
50026
d9871e5ea0e1 importing term with schematic type variables properly before passing it to the tactic in the set_comprehension_pointfree simproc
bulwahn
parents: 50025
diff changeset
   451
    val export = singleton (Variable.export ctxt' ctxt)
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   452
  in
50026
d9871e5ea0e1 importing term with schematic type variables properly before passing it to the tactic in the set_comprehension_pointfree simproc
bulwahn
parents: 50025
diff changeset
   453
    Option.map (export o post o unfold o mk_thm) (rewrite_term t'')
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   454
  end;
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   455
78801
42ae6e0ecfd4 tuned signature;
wenzelm
parents: 78800
diff changeset
   456
fun base_proc ctxt redex =
48122
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   457
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   458
    val set_compr = Thm.term_of redex
48122
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   459
  in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   460
    conv ctxt set_compr
48122
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   461
    |> Option.map (fn thm => thm RS @{thm eq_reflection})
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   462
  end;
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   463
78801
42ae6e0ecfd4 tuned signature;
wenzelm
parents: 78800
diff changeset
   464
fun proc ctxt redex =
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   465
  let
80655
be3325cbeb40 tuned: more antiquotations, avoid re-certification;
wenzelm
parents: 80642
diff changeset
   466
    val (f, set_compr) = Thm.dest_comb redex
be3325cbeb40 tuned: more antiquotations, avoid re-certification;
wenzelm
parents: 80642
diff changeset
   467
    val [A, B] = Thm.dest_ctyp (Thm.ctyp_of_cterm f)
be3325cbeb40 tuned: more antiquotations, avoid re-certification;
wenzelm
parents: 80642
diff changeset
   468
    val cong =
be3325cbeb40 tuned: more antiquotations, avoid re-certification;
wenzelm
parents: 80642
diff changeset
   469
      \<^instantiate>\<open>'a = A and 'b = B and f
be3325cbeb40 tuned: more antiquotations, avoid re-certification;
wenzelm
parents: 80642
diff changeset
   470
        in lemma (schematic) "x = y \<Longrightarrow> f x \<equiv> f y" by simp\<close>
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   471
  in
80655
be3325cbeb40 tuned: more antiquotations, avoid re-certification;
wenzelm
parents: 80642
diff changeset
   472
    conv ctxt (Thm.term_of set_compr)
be3325cbeb40 tuned: more antiquotations, avoid re-certification;
wenzelm
parents: 80642
diff changeset
   473
    |> Option.map (fn thm => thm RS cong)
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   474
  end;
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   475
78801
42ae6e0ecfd4 tuned signature;
wenzelm
parents: 78800
diff changeset
   476
fun code_proc ctxt redex =
48122
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   477
  let
50033
c78f9cddc907 rewriting with the simpset that is passed to the simproc
bulwahn
parents: 50032
diff changeset
   478
    fun unfold_conv thms =
c78f9cddc907 rewriting with the simpset that is passed to the simproc
bulwahn
parents: 50032
diff changeset
   479
      Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51315
diff changeset
   480
        (empty_simpset ctxt addsimps thms)
50033
c78f9cddc907 rewriting with the simpset that is passed to the simproc
bulwahn
parents: 50032
diff changeset
   481
    val prep_thm = unfold_conv @{thms eq_equal[symmetric]} redex
48122
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   482
  in
78801
42ae6e0ecfd4 tuned signature;
wenzelm
parents: 78800
diff changeset
   483
    case base_proc ctxt (Thm.rhs_of prep_thm) of
48122
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   484
      SOME rewr_thm => SOME (transitive_thm OF [transitive_thm OF [prep_thm, rewr_thm],
50033
c78f9cddc907 rewriting with the simpset that is passed to the simproc
bulwahn
parents: 50032
diff changeset
   485
        unfold_conv @{thms eq_equal} (Thm.rhs_of rewr_thm)])
48122
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   486
    | NONE => NONE
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   487
  end;
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   488
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   489
end;