src/HOL/Tools/set_comprehension_pointfree.ML
author bulwahn
Wed, 10 Oct 2012 10:48:55 +0200
changeset 49768 3ecfba7e731d
parent 49765 b9eb9c2b87c7
child 49831 b28dbb7a45d9
permissions -rw-r--r--
adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
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
48128
bf172a5929bb merged, resolving conflict with 87c831e30f0a;
wenzelm
parents: 48122 48124
diff changeset
    10
  val code_simproc : simpset -> cterm -> thm option
48124
87c831e30f0a ignore morphism more explicitly;
wenzelm
parents: 48109
diff changeset
    11
  val simproc : simpset -> cterm -> thm option
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    12
  val rewrite_term : term -> term option
48122
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
    13
  (* FIXME: function conv is not a conversion, i.e. of type cterm -> thm, MAYBE rename *)
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    14
  val conv : Proof.context -> term -> thm option
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    15
end
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    16
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    17
structure Set_Comprehension_Pointfree : SET_COMPREHENSION_POINTFREE =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    18
struct
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
(* syntactic operations *)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    21
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    22
fun mk_inf (t1, t2) =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    23
  let
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    24
    val T = fastype_of t1
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    25
  in
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    26
    Const (@{const_name Lattices.inf_class.inf}, T --> T --> T) $ t1 $ t2
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    27
  end
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    28
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
    29
fun mk_sup (t1, t2) =
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    30
  let
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
    val T = fastype_of t1
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    32
  in
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    33
    Const (@{const_name Lattices.sup_class.sup}, T --> T --> T) $ t1 $ t2
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    34
  end
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    35
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    36
fun mk_Compl t =
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    37
  let
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    38
    val T = fastype_of t
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    39
  in
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    40
    Const (@{const_name "Groups.uminus_class.uminus"}, T --> T) $ t
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    41
  end
3ecfba7e731d adding necessary syntactic functions in set_comprehension_pointfree simproc as a first step to integrate an improved simproc
bulwahn
parents: 49765
diff changeset
    42
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    43
fun mk_image t1 t2 =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    44
  let
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    45
    val T as Type (@{type_name fun}, [_ , R]) = fastype_of t1
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    46
  in
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    47
    Const (@{const_name image},
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    48
      T --> fastype_of t2 --> HOLogic.mk_setT R) $ t1 $ t2
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    49
  end;
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    50
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    51
fun mk_sigma (t1, t2) =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    52
  let
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    53
    val T1 = fastype_of t1
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    54
    val T2 = fastype_of t2
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    55
    val setT = HOLogic.dest_setT T1
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    56
    val resT = HOLogic.mk_setT (HOLogic.mk_prodT (setT, HOLogic.dest_setT T2))
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    57
  in
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    58
    Const (@{const_name Sigma},
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    59
      T1 --> (setT --> T2) --> resT) $ t1 $ absdummy setT t2
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    60
  end;
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    61
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    62
fun dest_Bound (Bound x) = x
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    63
  | dest_Bound t = raise TERM("dest_Bound", [t]);
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    64
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    65
fun dest_Collect (Const (@{const_name Collect}, _) $ Abs (_, _, t)) = t
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    66
  | 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
    67
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    68
(* Copied from predicate_compile_aux.ML *)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    69
fun strip_ex (Const (@{const_name Ex}, _) $ Abs (x, T, t)) =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    70
  let
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    71
    val (xTs, t') = strip_ex t
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    72
  in
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    73
    ((x, T) :: xTs, t')
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    74
  end
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    75
  | strip_ex t = ([], t)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    76
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    77
fun list_tupled_abs [] f = f
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    78
  | list_tupled_abs [(n, T)] f = (Abs (n, T, f))
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    79
  | list_tupled_abs ((n, T)::v::vs) f =
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    80
      HOLogic.mk_split (Abs (n, T, list_tupled_abs (v::vs) f))
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    81
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    82
fun mk_pointfree_expr t =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    83
  let
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    84
    val (vs, t'') = strip_ex (dest_Collect t)
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
    85
    val conjs = HOLogic.dest_conj t''
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
    86
    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
    87
      the_default false o (try (fn eq => fst (HOLogic.dest_eq eq) = Bound (length vs)))
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
    88
    val SOME eq = find_first is_the_eq conjs
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
    89
    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
    90
    val conjs' = filter_out (fn t => eq = t) conjs
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
    91
    val mems = map (apfst dest_Bound o HOLogic.dest_mem) conjs'
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    92
    val grouped_mems = AList.group (op =) mems
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    93
    fun mk_grouped_unions (i, T) =
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    94
      case AList.lookup (op =) grouped_mems i of
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    95
        SOME ts => foldr1 mk_inf ts
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    96
      | NONE => HOLogic.mk_UNIV T
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    97
    val complete_sets = map mk_grouped_unions ((length vs - 1) downto 0 ~~ map snd vs)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    98
  in
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    99
    mk_image (list_tupled_abs vs f) (foldr1 mk_sigma complete_sets)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   100
  end;
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   101
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   102
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
   103
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   104
(* proof tactic *)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   105
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   106
(* Tactic works for arbitrary number of m : S conjuncts *)
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   107
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   108
val dest_Collect_tac = dtac @{thm iffD1[OF mem_Collect_eq]}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   109
  THEN' (REPEAT_DETERM o (eresolve_tac @{thms exE conjE}))
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   110
  THEN' hyp_subst_tac;
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   111
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   112
val intro_image_Sigma_tac = rtac @{thm image_eqI}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   113
    THEN' (REPEAT_DETERM1 o
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   114
      (rtac @{thm refl}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   115
      ORELSE' rtac
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   116
        @{thm arg_cong2[OF refl, where f="op =", OF prod.cases, THEN iffD2]}));
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   117
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   118
val dest_image_Sigma_tac = etac @{thm imageE}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   119
  THEN' (TRY o REPEAT_DETERM1 o Splitter.split_asm_tac @{thms prod.split_asm})
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   120
  THEN' hyp_subst_tac
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   121
  THEN' (TRY o REPEAT_DETERM1 o
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   122
    (etac @{thm conjE} ORELSE' dtac @{thm iffD1[OF mem_Sigma_iff]}));
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   123
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   124
val intro_Collect_tac = rtac @{thm iffD2[OF mem_Collect_eq]}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   125
  THEN' REPEAT_DETERM1 o resolve_tac @{thms exI}
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
   126
  THEN' (TRY o REPEAT_ALL_NEW (rtac @{thm conjI}))
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
   127
  THEN' (K (ALLGOALS (TRY o ((TRY o hyp_subst_tac) THEN' rtac @{thm refl}))))
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   128
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   129
val tac =
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   130
  let
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   131
    val subset_tac1 = rtac @{thm subsetI}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   132
      THEN' dest_Collect_tac
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   133
      THEN' intro_image_Sigma_tac
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   134
      THEN' (REPEAT_DETERM1 o
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   135
        (rtac @{thm SigmaI}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   136
        ORELSE' rtac @{thm UNIV_I}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   137
        ORELSE' rtac @{thm IntI}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   138
        ORELSE' atac));
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   139
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   140
    val subset_tac2 = rtac @{thm subsetI}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   141
      THEN' dest_image_Sigma_tac
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   142
      THEN' intro_Collect_tac
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
   143
      THEN' REPEAT_DETERM o (eresolve_tac @{thms IntD1 IntD2} ORELSE' atac);
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   144
  in
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   145
    rtac @{thm subset_antisym} THEN' subset_tac1 THEN' subset_tac2
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   146
  end;
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   147
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   148
fun conv ctxt t =
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   149
  let
49765
b9eb9c2b87c7 unfolding bounded existential quantifiers as first step in the set_comprehension_pointfree simproc
bulwahn
parents: 49763
diff changeset
   150
    val ct = cterm_of (Proof_Context.theory_of ctxt) t
b9eb9c2b87c7 unfolding bounded existential quantifiers as first step in the set_comprehension_pointfree simproc
bulwahn
parents: 49763
diff changeset
   151
    val Bex_def = mk_meta_eq @{thm Bex_def}
b9eb9c2b87c7 unfolding bounded existential quantifiers as first step in the set_comprehension_pointfree simproc
bulwahn
parents: 49763
diff changeset
   152
    val unfold_eq = Conv.bottom_conv (K (Conv.try_conv (Conv.rewr_conv Bex_def))) ctxt ct
b9eb9c2b87c7 unfolding bounded existential quantifiers as first step in the set_comprehension_pointfree simproc
bulwahn
parents: 49763
diff changeset
   153
    val t' = term_of (Thm.rhs_of unfold_eq) 
b9eb9c2b87c7 unfolding bounded existential quantifiers as first step in the set_comprehension_pointfree simproc
bulwahn
parents: 49763
diff changeset
   154
    fun mk_thm t'' = Goal.prove ctxt [] []
b9eb9c2b87c7 unfolding bounded existential quantifiers as first step in the set_comprehension_pointfree simproc
bulwahn
parents: 49763
diff changeset
   155
      (HOLogic.mk_Trueprop (HOLogic.mk_eq (t', t''))) (K (tac 1))
b9eb9c2b87c7 unfolding bounded existential quantifiers as first step in the set_comprehension_pointfree simproc
bulwahn
parents: 49763
diff changeset
   156
    fun unfold th = th RS ((unfold_eq RS meta_eq_to_obj_eq) RS @{thm trans})
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   157
  in
49765
b9eb9c2b87c7 unfolding bounded existential quantifiers as first step in the set_comprehension_pointfree simproc
bulwahn
parents: 49763
diff changeset
   158
    Option.map (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
   159
  end;
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   160
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   161
(* simproc *)
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   162
48128
bf172a5929bb merged, resolving conflict with 87c831e30f0a;
wenzelm
parents: 48122 48124
diff changeset
   163
fun base_simproc ss redex =
48122
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   164
  let
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   165
    val ctxt = Simplifier.the_context ss
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   166
    val set_compr = term_of redex
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   167
  in
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   168
    conv ctxt set_compr
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   169
    |> 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
   170
  end;
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   171
49763
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   172
fun instantiate_arg_cong ctxt pred =
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   173
  let
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   174
    val certify = cterm_of (Proof_Context.theory_of ctxt)
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   175
    val arg_cong = @{thm arg_cong}
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   176
    val f $ _ = fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (concl_of arg_cong)))
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   177
  in
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   178
    cterm_instantiate [(certify f, certify pred)] arg_cong
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   179
  end;
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   180
48124
87c831e30f0a ignore morphism more explicitly;
wenzelm
parents: 48109
diff changeset
   181
fun simproc ss redex =
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   182
  let
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   183
    val ctxt = Simplifier.the_context ss
49763
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   184
    val pred $ set_compr = term_of redex
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   185
    val arg_cong' = instantiate_arg_cong ctxt pred
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   186
  in
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   187
    conv ctxt set_compr
49763
bed063d0c526 generalizing set_comprehension_pointfree simproc to work for arbitrary predicates (and not just the finite predicate)
bulwahn
parents: 49761
diff changeset
   188
    |> Option.map (fn thm => thm RS arg_cong' RS @{thm eq_reflection})
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   189
  end;
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   190
48128
bf172a5929bb merged, resolving conflict with 87c831e30f0a;
wenzelm
parents: 48122 48124
diff changeset
   191
fun code_simproc ss redex =
48122
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   192
  let
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   193
    val prep_thm = Raw_Simplifier.rewrite false @{thms eq_equal[symmetric]} redex
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   194
  in
48128
bf172a5929bb merged, resolving conflict with 87c831e30f0a;
wenzelm
parents: 48122 48124
diff changeset
   195
    case base_simproc ss (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
   196
      SOME rewr_thm => SOME (transitive_thm OF [transitive_thm OF [prep_thm, rewr_thm],
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   197
        Raw_Simplifier.rewrite false @{thms eq_equal} (Thm.rhs_of rewr_thm)])
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   198
    | NONE => NONE
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   199
  end;
f479f36ed2ff adding set comprehension simproc to code generation's preprocessing to generate code for some set comprehensions;
bulwahn
parents: 48109
diff changeset
   200
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
   201
end;
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
   202