author | wenzelm |
Thu, 01 Sep 2016 21:28:46 +0200 | |
changeset 63763 | 0f61ea70d384 |
parent 63167 | 0909deb8059b |
child 66453 | cc19f7ca2ed6 |
permissions | -rw-r--r-- |
33027 | 1 |
(* Title: HOL/Metis_Examples/Abstraction.thy |
43197 | 2 |
Author: Lawrence C. Paulson, Cambridge University Computer Laboratory |
41144 | 3 |
Author: Jasmin Blanchette, TU Muenchen |
23449 | 4 |
|
43197 | 5 |
Example featuring Metis's support for lambda-abstractions. |
23449 | 6 |
*) |
7 |
||
63167 | 8 |
section \<open>Example Featuring Metis's Support for Lambda-Abstractions\<close> |
43197 | 9 |
|
27368 | 10 |
theory Abstraction |
45572 | 11 |
imports "~~/src/HOL/Library/FuncSet" |
23449 | 12 |
begin |
13 |
||
45562 | 14 |
(* For Christoph Benzmüller *) |
15 |
lemma "x < 1 \<and> ((op =) = (op =)) \<Longrightarrow> ((op =) = (op =)) \<and> x < (2::nat)" |
|
16 |
by (metis nat_1_add_1 trans_less_add2) |
|
23449 | 17 |
|
45572 | 18 |
lemma "(op = ) = (\<lambda>x y. y = x)" |
45562 | 19 |
by metis |
23449 | 20 |
|
21 |
consts |
|
22 |
monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool" |
|
23 |
pset :: "'a set => 'a set" |
|
24 |
order :: "'a set => ('a * 'a) set" |
|
25 |
||
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
26 |
lemma "a \<in> {x. P x} \<Longrightarrow> P a" |
36566 | 27 |
proof - |
28 |
assume "a \<in> {x. P x}" |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
29 |
thus "P a" by (metis mem_Collect_eq) |
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
30 |
qed |
23449 | 31 |
|
45572 | 32 |
lemma Collect_triv: "a \<in> {x. P x} \<Longrightarrow> P a" |
23756 | 33 |
by (metis mem_Collect_eq) |
23449 | 34 |
|
45572 | 35 |
lemma "a \<in> {x. P x --> Q x} \<Longrightarrow> a \<in> {x. P x} \<Longrightarrow> a \<in> {x. Q x}" |
45562 | 36 |
by (metis Collect_imp_eq ComplD UnE) |
23449 | 37 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
38 |
lemma "(a, b) \<in> Sigma A B \<Longrightarrow> a \<in> A \<and> b \<in> B a" |
36566 | 39 |
proof - |
40 |
assume A1: "(a, b) \<in> Sigma A B" |
|
41 |
hence F1: "b \<in> B a" by (metis mem_Sigma_iff) |
|
42 |
have F2: "a \<in> A" by (metis A1 mem_Sigma_iff) |
|
43 |
have "b \<in> B a" by (metis F1) |
|
44 |
thus "a \<in> A \<and> b \<in> B a" by (metis F2) |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
45 |
qed |
23449 | 46 |
|
45572 | 47 |
lemma Sigma_triv: "(a, b) \<in> Sigma A B \<Longrightarrow> a \<in> A & b \<in> B a" |
23449 | 48 |
by (metis SigmaD1 SigmaD2) |
49 |
||
36566 | 50 |
lemma "(a, b) \<in> (SIGMA x:A. {y. x = f y}) \<Longrightarrow> a \<in> A \<and> a = f b" |
46364 | 51 |
by (metis (full_types, lifting) CollectD SigmaD1 SigmaD2) |
24827 | 52 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
53 |
lemma "(a, b) \<in> (SIGMA x:A. {y. x = f y}) \<Longrightarrow> a \<in> A \<and> a = f b" |
36566 | 54 |
proof - |
55 |
assume A1: "(a, b) \<in> (SIGMA x:A. {y. x = f y})" |
|
56 |
hence F1: "a \<in> A" by (metis mem_Sigma_iff) |
|
57 |
have "b \<in> {R. a = f R}" by (metis A1 mem_Sigma_iff) |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
58 |
hence "a = f b" by (metis (full_types) mem_Collect_eq) |
36566 | 59 |
thus "a \<in> A \<and> a = f b" by (metis F1) |
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
60 |
qed |
23449 | 61 |
|
45572 | 62 |
lemma "(cl, f) \<in> CLF \<Longrightarrow> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) \<Longrightarrow> f \<in> pset cl" |
24827 | 63 |
by (metis Collect_mem_eq SigmaD2) |
23449 | 64 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
65 |
lemma "(cl, f) \<in> CLF \<Longrightarrow> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) \<Longrightarrow> f \<in> pset cl" |
36566 | 66 |
proof - |
67 |
assume A1: "(cl, f) \<in> CLF" |
|
68 |
assume A2: "CLF = (SIGMA cl:CL. {f. f \<in> pset cl})" |
|
69 |
have "\<forall>v u. (u, v) \<in> CLF \<longrightarrow> v \<in> {R. R \<in> pset u}" by (metis A2 mem_Sigma_iff) |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
70 |
hence "\<forall>v u. (u, v) \<in> CLF \<longrightarrow> v \<in> pset u" by (metis mem_Collect_eq) |
45572 | 71 |
thus "f \<in> pset cl" by (metis A1) |
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
72 |
qed |
23449 | 73 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
74 |
lemma |
45572 | 75 |
"(cl, f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) \<Longrightarrow> |
76 |
f \<in> pset cl \<rightarrow> pset cl" |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
77 |
by (metis (no_types) Collect_mem_eq Sigma_triv) |
45562 | 78 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
79 |
lemma |
45572 | 80 |
"(cl, f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) \<Longrightarrow> |
81 |
f \<in> pset cl \<rightarrow> pset cl" |
|
36566 | 82 |
proof - |
83 |
assume A1: "(cl, f) \<in> (SIGMA cl:CL. {f. f \<in> pset cl \<rightarrow> pset cl})" |
|
84 |
have "f \<in> {R. R \<in> pset cl \<rightarrow> pset cl}" using A1 by simp |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
85 |
thus "f \<in> pset cl \<rightarrow> pset cl" by (metis mem_Collect_eq) |
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
86 |
qed |
23449 | 87 |
|
88 |
lemma |
|
45572 | 89 |
"(cl, f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) \<Longrightarrow> |
90 |
f \<in> pset cl \<inter> cl" |
|
45562 | 91 |
by (metis (no_types) Collect_conj_eq Int_def Sigma_triv inf_idem) |
92 |
||
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
93 |
lemma |
45572 | 94 |
"(cl, f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) \<Longrightarrow> |
95 |
f \<in> pset cl \<inter> cl" |
|
36566 | 96 |
proof - |
97 |
assume A1: "(cl, f) \<in> (SIGMA cl:CL. {f. f \<in> pset cl \<inter> cl})" |
|
98 |
have "f \<in> {R. R \<in> pset cl \<inter> cl}" using A1 by simp |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
99 |
hence "f \<in> Id_on cl `` pset cl" by (metis Int_commute Image_Id_on mem_Collect_eq) |
36566 | 100 |
hence "f \<in> cl \<inter> pset cl" by (metis Image_Id_on) |
101 |
thus "f \<in> pset cl \<inter> cl" by (metis Int_commute) |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
102 |
qed |
24827 | 103 |
|
23449 | 104 |
lemma |
45572 | 105 |
"(cl, f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) \<Longrightarrow> |
23449 | 106 |
(f \<in> pset cl \<rightarrow> pset cl) & (monotone f (pset cl) (order cl))" |
107 |
by auto |
|
108 |
||
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
109 |
lemma |
45572 | 110 |
"(cl, f) \<in> CLF \<Longrightarrow> |
111 |
CLF \<subseteq> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) \<Longrightarrow> |
|
112 |
f \<in> pset cl \<inter> cl" |
|
46364 | 113 |
by (metis (lifting) CollectD Sigma_triv subsetD) |
36566 | 114 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
115 |
lemma |
45572 | 116 |
"(cl, f) \<in> CLF \<Longrightarrow> |
117 |
CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) \<Longrightarrow> |
|
118 |
f \<in> pset cl \<inter> cl" |
|
46364 | 119 |
by (metis (lifting) CollectD Sigma_triv) |
36566 | 120 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
121 |
lemma |
45572 | 122 |
"(cl, f) \<in> CLF \<Longrightarrow> |
123 |
CLF \<subseteq> (SIGMA cl': CL. {f. f \<in> pset cl' \<rightarrow> pset cl'}) \<Longrightarrow> |
|
124 |
f \<in> pset cl \<rightarrow> pset cl" |
|
46364 | 125 |
by (metis (lifting) CollectD Sigma_triv subsetD) |
45572 | 126 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
127 |
lemma |
45572 | 128 |
"(cl, f) \<in> CLF \<Longrightarrow> |
129 |
CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) \<Longrightarrow> |
|
130 |
f \<in> pset cl \<rightarrow> pset cl" |
|
46364 | 131 |
by (metis (lifting) CollectD Sigma_triv) |
45572 | 132 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
133 |
lemma |
45572 | 134 |
"(cl, f) \<in> CLF \<Longrightarrow> |
135 |
CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) \<Longrightarrow> |
|
136 |
(f \<in> pset cl \<rightarrow> pset cl) & (monotone f (pset cl) (order cl))" |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
137 |
by auto |
23449 | 138 |
|
45572 | 139 |
lemma "map (\<lambda>x. (f x, g x)) xs = zip (map f xs) (map g xs)" |
23449 | 140 |
apply (induct xs) |
55465 | 141 |
apply (metis list.map(1) zip_Nil) |
36566 | 142 |
by auto |
23449 | 143 |
|
45572 | 144 |
lemma |
61384 | 145 |
"map (\<lambda>w. (w \<rightarrow> w, w \<times> w)) xs = |
146 |
zip (map (\<lambda>w. w \<rightarrow> w) xs) (map (\<lambda>w. w \<times> w) xs)" |
|
45572 | 147 |
apply (induct xs) |
55465 | 148 |
apply (metis list.map(1) zip_Nil) |
45572 | 149 |
by auto |
150 |
||
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
151 |
lemma "(\<lambda>x. Suc (f x)) ` {x. even x} \<subseteq> A \<Longrightarrow> \<forall>x. even x --> Suc (f x) \<in> A" |
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
152 |
by (metis mem_Collect_eq image_eqI subsetD) |
23449 | 153 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
154 |
lemma |
45572 | 155 |
"(\<lambda>x. f (f x)) ` ((\<lambda>x. Suc(f x)) ` {x. even x}) \<subseteq> A \<Longrightarrow> |
156 |
(\<forall>x. even x --> f (f (Suc(f x))) \<in> A)" |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
157 |
by (metis mem_Collect_eq imageI set_rev_mp) |
23449 | 158 |
|
46076
a109eb27f54f
ported a dozen of proofs to the "set" type constructor
blanchet
parents:
45972
diff
changeset
|
159 |
lemma "f \<in> (\<lambda>u v. b \<times> u \<times> v) ` A \<Longrightarrow> \<forall>u v. P (b \<times> u \<times> v) \<Longrightarrow> P(f y)" |
46364 | 160 |
by (metis (lifting) imageE) |
23449 | 161 |
|
45572 | 162 |
lemma image_TimesA: "(\<lambda>(x, y). (f x, g y)) ` (A \<times> B) = (f ` A) \<times> (g ` B)" |
55932 | 163 |
by (metis map_prod_def map_prod_surj_on) |
23449 | 164 |
|
165 |
lemma image_TimesB: |
|
45572 | 166 |
"(\<lambda>(x, y, z). (f x, g y, h z)) ` (A \<times> B \<times> C) = (f ` A) \<times> (g ` B) \<times> (h ` C)" |
23449 | 167 |
by force |
168 |
||
169 |
lemma image_TimesC: |
|
45572 | 170 |
"(\<lambda>(x, y). (x \<rightarrow> x, y \<times> y)) ` (A \<times> B) = |
171 |
((\<lambda>x. x \<rightarrow> x) ` A) \<times> ((\<lambda>y. y \<times> y) ` B)" |
|
45562 | 172 |
by (metis image_TimesA) |
23449 | 173 |
|
174 |
end |