src/HOL/ex/Set_Comprehension_Pointfree_Tests.thy
author Christian Sternagel
Thu, 30 Aug 2012 15:44:03 +0900
changeset 49093 fdc301f592c4
parent 48174 eb72f99737be
child 49762 b5e355c41de3
permissions -rw-r--r--
forgot to add lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     1
(*  Title:      HOL/ex/Set_Comprehension_Pointfree_Tests.thy
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
     2
    Author:     Lukas Bulwahn, Rafal Kolanski
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     3
    Copyright   2012 TU Muenchen
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
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     6
header {* Tests for the set comprehension to pointfree simproc *}
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
theory Set_Comprehension_Pointfree_Tests
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
     9
imports Main
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    10
begin
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    11
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    12
lemma
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    13
  "finite {p. EX x. p = (x, x)}"
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    14
  apply simp
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    15
  oops
48049
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
lemma
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    18
  "finite {f a b| a b. a : A \<and> b : B}"
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    19
  apply simp
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    20
  oops
48049
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
lemma
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    23
  "finite {f a b| a b. a : A \<and> a : A' \<and> b : B}"
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    24
  apply simp
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    25
  oops
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    26
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    27
lemma
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    28
  "finite {f a b| a b. a : A \<and> b : B \<and> b : B'}"
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    29
  apply simp
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    30
  oops
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    31
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    32
lemma
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    33
  "finite {f a b c| a b c. a : A \<and> b : B \<and> c : C}"
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    34
  apply simp
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    35
  oops
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    36
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    37
lemma
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    38
  "finite {f a b c d| a b c d. a : A \<and> b : B \<and> c : C \<and> d : D}"
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    39
  apply simp
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    40
  oops
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    41
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    42
lemma
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    43
  "finite {f a b c d e | a b c d e. a : A \<and> b : B \<and> c : C \<and> d : D \<and> e : E}"
48108
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    44
  apply simp
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    45
  oops
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    46
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    47
lemma (* check arbitrary ordering *)
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    48
  "finite {f a d c b e | e b c d a. b : B \<and> a : A \<and> e : E' \<and> c : C \<and> d : D \<and> e : E \<and> b : B'}"
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    49
  apply simp
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    50
  oops
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    51
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    52
lemma
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    53
  "\<lbrakk> finite A ; finite B ; finite C ; finite D \<rbrakk>
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    54
  \<Longrightarrow> finite ({f a b c d| a b c d. a : A \<and> b : B \<and> c : C \<and> d : D})"
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    55
  by simp
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    56
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    57
lemma
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    58
  "finite ((\<lambda>(a,b,c,d). f a b c d) ` (A \<times> B \<times> C \<times> D))
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    59
  \<Longrightarrow> finite ({f a b c d| a b c d. a : A \<and> b : B \<and> c : C \<and> d : D})"
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    60
  by simp
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    61
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    62
schematic_lemma (* check interaction with schematics *)
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    63
  "finite {x :: ?'A \<Rightarrow> ?'B \<Rightarrow> bool. \<exists>a b. x = Pair_Rep a b}
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    64
   = finite ((\<lambda>(a:: ?'A, b :: ?'B). Pair_Rep a b) ` (UNIV \<times> UNIV))"
f93433b451b8 Improved tactic for rewriting set comprehensions into pointfree form.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48055
diff changeset
    65
  by simp
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    66
48174
eb72f99737be adding a challenging example in the examples file
bulwahn
parents: 48109
diff changeset
    67
lemma
eb72f99737be adding a challenging example in the examples file
bulwahn
parents: 48109
diff changeset
    68
  assumes "finite S" shows "finite {(a,b,c,d). ([a, b], [c, d]) : S}"
eb72f99737be adding a challenging example in the examples file
bulwahn
parents: 48109
diff changeset
    69
proof -
eb72f99737be adding a challenging example in the examples file
bulwahn
parents: 48109
diff changeset
    70
  have eq: "{(a,b,c,d). ([a, b], [c, d]) : S} = ((%(a, b, c, d). ([a, b], [c, d])) -` S)"
eb72f99737be adding a challenging example in the examples file
bulwahn
parents: 48109
diff changeset
    71
   unfolding vimage_def by (auto split: prod.split)  (* to be proved with the simproc *)
eb72f99737be adding a challenging example in the examples file
bulwahn
parents: 48109
diff changeset
    72
  from `finite S` show ?thesis
eb72f99737be adding a challenging example in the examples file
bulwahn
parents: 48109
diff changeset
    73
    unfolding eq by (auto intro!: finite_vimageI simp add: inj_on_def)
eb72f99737be adding a challenging example in the examples file
bulwahn
parents: 48109
diff changeset
    74
    (* to be automated with further rules and automation *)
eb72f99737be adding a challenging example in the examples file
bulwahn
parents: 48109
diff changeset
    75
qed
eb72f99737be adding a challenging example in the examples file
bulwahn
parents: 48109
diff changeset
    76
48049
d862b0d56c49 adding incompleted simproc to rewrite set comprehensions into pointfree expressions on sets
bulwahn
parents:
diff changeset
    77
end