src/HOL/ex/Set_Comprehension_Pointfree_Tests.thy
author traytel
Fri, 29 Nov 2013 08:26:45 +0100
changeset 55984 31afce809794
parent 50962 29cd291bfea6
child 56276 4587de627cd8
permissions -rw-r--r--
set_comprehension_pointfree simproc causes to many surprises if enabled by default
bulwahn@49064
     1
(*  Title:      HOL/ex/Set_Comprehension_Pointfree_Tests.thy
rafal@49123
     2
    Author:     Lukas Bulwahn, Rafal Kolanski
bulwahn@49064
     3
    Copyright   2012 TU Muenchen
bulwahn@49064
     4
*)
bulwahn@49064
     5
bulwahn@49064
     6
header {* Tests for the set comprehension to pointfree simproc *}
bulwahn@49064
     7
bulwahn@49064
     8
theory Set_Comprehension_Pointfree_Tests
bulwahn@49064
     9
imports Main
bulwahn@49064
    10
begin
bulwahn@49064
    11
traytel@55984
    12
declare [[simproc add: finite_Collect]]
traytel@55984
    13
rafal@49123
    14
lemma
bulwahn@50866
    15
  "finite (UNIV::'a set) ==> finite {p. EX x::'a. p = (x, x)}"
bulwahn@50866
    16
  by simp
bulwahn@49064
    17
bulwahn@49064
    18
lemma
bulwahn@50866
    19
  "finite A ==> finite B ==> finite {f a b| a b. a : A \<and> b : B}"
bulwahn@50866
    20
  by simp
bulwahn@50866
    21
  
bulwahn@50866
    22
lemma
bulwahn@50866
    23
  "finite B ==> finite A' ==> finite {f a b| a b. a : A \<and> a : A' \<and> b : B}"
bulwahn@50866
    24
  by simp
bulwahn@49064
    25
bulwahn@49064
    26
lemma
bulwahn@50866
    27
  "finite A ==> finite B ==> finite {f a b| a b. a : A \<and> b : B \<and> b : B'}"
bulwahn@50866
    28
  by simp
bulwahn@49064
    29
bulwahn@49064
    30
lemma
bulwahn@50866
    31
  "finite A ==> finite B ==> finite C ==> finite {f a b c| a b c. a : A \<and> b : B \<and> c : C}"
bulwahn@50866
    32
  by simp
bulwahn@49064
    33
bulwahn@49064
    34
lemma
bulwahn@50866
    35
  "finite A ==> finite B ==> finite C ==> finite D ==>
bulwahn@50866
    36
     finite {f a b c d| a b c d. a : A \<and> b : B \<and> c : C \<and> d : D}"
bulwahn@50866
    37
  by simp
bulwahn@49064
    38
bulwahn@49064
    39
lemma
bulwahn@50866
    40
  "finite A ==> finite B ==> finite C ==> finite D ==> finite E ==>
bulwahn@50866
    41
    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}"
bulwahn@50866
    42
  by simp
rafal@49123
    43
bulwahn@50914
    44
lemma
bulwahn@50866
    45
  "finite A ==> finite B ==> finite C ==> finite D ==> finite E \<Longrightarrow>
bulwahn@50866
    46
    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'}"
bulwahn@50866
    47
  by simp
rafal@49123
    48
rafal@49123
    49
lemma
rafal@49123
    50
  "\<lbrakk> finite A ; finite B ; finite C ; finite D \<rbrakk>
rafal@49123
    51
  \<Longrightarrow> finite ({f a b c d| a b c d. a : A \<and> b : B \<and> c : C \<and> d : D})"
rafal@49123
    52
  by simp
rafal@49123
    53
rafal@49123
    54
lemma
rafal@49123
    55
  "finite ((\<lambda>(a,b,c,d). f a b c d) ` (A \<times> B \<times> C \<times> D))
rafal@49123
    56
  \<Longrightarrow> finite ({f a b c d| a b c d. a : A \<and> b : B \<and> c : C \<and> d : D})"
rafal@49123
    57
  by simp
rafal@49123
    58
bulwahn@50777
    59
lemma
bulwahn@50781
    60
  "finite S ==> finite {s'. EX s:S. s' = f a e s}"
bulwahn@50781
    61
  by simp
bulwahn@50777
    62
bulwahn@50866
    63
lemma
bulwahn@50866
    64
  "finite A ==> finite B ==> finite {f a b| a b. a : A \<and> b : B \<and> a \<notin> Z}"
bulwahn@50866
    65
  by simp
bulwahn@50866
    66
bulwahn@50866
    67
lemma
bulwahn@50866
    68
  "finite A ==> finite B ==> finite R ==> finite {f a b x y| a b x y. a : A \<and> b : B \<and> (x,y) \<in> R}"
bulwahn@50866
    69
by simp
bulwahn@50866
    70
bulwahn@50866
    71
lemma
bulwahn@50866
    72
  "finite A ==> finite B ==> finite R ==> finite {f a b x y| a b x y. a : A \<and> (x,y) \<in> R \<and> b : B}"
bulwahn@50866
    73
by simp
bulwahn@50866
    74
bulwahn@50866
    75
lemma
bulwahn@50866
    76
  "finite A ==> finite B ==> finite R ==> finite {f a (x, b) y| y b x a. a : A \<and> (x,y) \<in> R \<and> b : B}"
bulwahn@50866
    77
by simp
bulwahn@50866
    78
bulwahn@50866
    79
lemma
bulwahn@50866
    80
  "finite A ==> finite AA ==> finite B ==> finite {f a b| a b. (a : A \<or> a : AA) \<and> b : B \<and> a \<notin> Z}"
bulwahn@50866
    81
by simp
bulwahn@50866
    82
bulwahn@50866
    83
lemma
bulwahn@50866
    84
  "finite A1 ==> finite A2 ==> finite A3 ==> finite A4 ==> finite A5 ==> finite B ==>
bulwahn@50866
    85
     finite {f a b c | a b c. ((a : A1 \<and> a : A2) \<or> (a : A3 \<and> (a : A4 \<or> a : A5))) \<and> b : B \<and> a \<notin> Z}"
bulwahn@50866
    86
apply simp
bulwahn@50866
    87
oops
bulwahn@50866
    88
bulwahn@50868
    89
lemma "finite B ==> finite {c. EX x. x : B & c = a * x}"
bulwahn@50868
    90
by simp
bulwahn@50868
    91
bulwahn@50868
    92
lemma
bulwahn@50868
    93
  "finite A ==> finite B ==> finite {f a * g b |a b. a : A & b : B}"
bulwahn@50868
    94
by simp
bulwahn@50868
    95
bulwahn@50891
    96
lemma
bulwahn@50891
    97
  "finite S ==> inj (%(x, y). g x y) ==> finite {f x y| x y. g x y : S}"
bulwahn@50891
    98
  by (auto intro: finite_vimageI)
bulwahn@50891
    99
bulwahn@50891
   100
lemma
bulwahn@50891
   101
  "finite A ==> finite S ==> inj (%(x, y). g x y) ==> finite {f x y z | x y z. g x y : S & z : A}"
bulwahn@50891
   102
  by (auto intro: finite_vimageI)
bulwahn@50891
   103
bulwahn@50891
   104
lemma
bulwahn@50891
   105
  "finite S ==> finite A ==> inj (%(x, y). g x y) ==> inj (%(x, y). h x y)
bulwahn@50891
   106
    ==> finite {f a b c d | a b c d. g a c : S & h b d : A}"
bulwahn@50891
   107
  by (auto intro: finite_vimageI)
bulwahn@50891
   108
bulwahn@50914
   109
lemma
bulwahn@50914
   110
  assumes "finite S" shows "finite {(a,b,c,d). ([a, b], [c, d]) : S}"
bulwahn@50914
   111
using assms by (auto intro!: finite_vimageI simp add: inj_on_def)
bulwahn@50914
   112
  (* injectivity to be automated with further rules and automation *)
bulwahn@50891
   113
rafal@49123
   114
schematic_lemma (* check interaction with schematics *)
rafal@49123
   115
  "finite {x :: ?'A \<Rightarrow> ?'B \<Rightarrow> bool. \<exists>a b. x = Pair_Rep a b}
bulwahn@50866
   116
   = finite ((\<lambda>(b :: ?'B, a:: ?'A). Pair_Rep a b) ` (UNIV \<times> UNIV))"
rafal@49123
   117
  by simp
bulwahn@49064
   118
traytel@55984
   119
declare [[simproc del: finite_Collect]]
traytel@55984
   120
bulwahn@49189
   121
bulwahn@50891
   122
section {* Testing simproc in code generation *}
bulwahn@50891
   123
bulwahn@50891
   124
definition union :: "nat set => nat set => nat set"
bulwahn@50891
   125
where
bulwahn@50891
   126
  "union A B = {x. x : A \<or> x : B}"
bulwahn@50891
   127
bulwahn@50891
   128
definition common_subsets :: "nat set => nat set => nat set set"
bulwahn@50891
   129
where
bulwahn@50891
   130
  "common_subsets S1 S2 = {S. S \<subseteq> S1 \<and> S \<subseteq> S2}"
bulwahn@50891
   131
bulwahn@50962
   132
definition products :: "nat set => nat set => nat set"
bulwahn@50962
   133
where
bulwahn@50962
   134
  "products A B = {c. EX a b. a : A & b : B & c = a * b}"
bulwahn@50962
   135
bulwahn@50962
   136
export_code products in Haskell file -
bulwahn@50962
   137
bulwahn@50962
   138
export_code union common_subsets products in Haskell file -
bulwahn@50891
   139
bulwahn@49064
   140
end