| author | blanchet |
| Fri, 18 Nov 2011 11:47:12 +0100 | |
| changeset 45563 | 94ebb64b0433 |
| parent 45562 | e8fab4786b3c |
| child 45572 | 08970468f99b |
| 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 |
||
| 43197 | 8 |
header {* Example Featuring Metis's Support for Lambda-Abstractions *}
|
9 |
||
| 27368 | 10 |
theory Abstraction |
|
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
41144
diff
changeset
|
11 |
imports Main "~~/src/HOL/Library/FuncSet" |
| 23449 | 12 |
begin |
13 |
||
|
42103
6066a35f6678
Metis examples use the new Skolemizer to test it
blanchet
parents:
41413
diff
changeset
|
14 |
declare [[metis_new_skolemizer]] |
|
6066a35f6678
Metis examples use the new Skolemizer to test it
blanchet
parents:
41413
diff
changeset
|
15 |
|
| 45562 | 16 |
(* For Christoph Benzmüller *) |
17 |
lemma "x < 1 \<and> ((op =) = (op =)) \<Longrightarrow> ((op =) = (op =)) \<and> x < (2::nat)" |
|
18 |
by (metis nat_1_add_1 trans_less_add2) |
|
| 23449 | 19 |
|
| 45562 | 20 |
lemma "(op = ) = (%x y. y = x)" |
21 |
by metis |
|
| 23449 | 22 |
|
23 |
consts |
|
24 |
monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool"
|
|
25 |
pset :: "'a set => 'a set" |
|
26 |
order :: "'a set => ('a * 'a) set"
|
|
27 |
||
28 |
lemma (*Collect_triv:*) "a \<in> {x. P x} ==> P a"
|
|
| 36566 | 29 |
proof - |
30 |
assume "a \<in> {x. P x}"
|
|
31 |
hence "a \<in> P" by (metis Collect_def) |
|
| 45562 | 32 |
thus "P a" by (metis mem_def) |
| 23449 | 33 |
qed |
34 |
||
35 |
lemma Collect_triv: "a \<in> {x. P x} ==> P a"
|
|
| 23756 | 36 |
by (metis mem_Collect_eq) |
| 23449 | 37 |
|
38 |
lemma "a \<in> {x. P x --> Q x} ==> a \<in> {x. P x} ==> a \<in> {x. Q x}"
|
|
| 45562 | 39 |
by (metis Collect_imp_eq ComplD UnE) |
| 23449 | 40 |
|
| 45562 | 41 |
lemma "(a, b) \<in> Sigma A B ==> a \<in> A & b \<in> B a" |
| 36566 | 42 |
proof - |
43 |
assume A1: "(a, b) \<in> Sigma A B" |
|
44 |
hence F1: "b \<in> B a" by (metis mem_Sigma_iff) |
|
45 |
have F2: "a \<in> A" by (metis A1 mem_Sigma_iff) |
|
46 |
have "b \<in> B a" by (metis F1) |
|
47 |
thus "a \<in> A \<and> b \<in> B a" by (metis F2) |
|
| 23449 | 48 |
qed |
49 |
||
50 |
lemma Sigma_triv: "(a,b) \<in> Sigma A B ==> a \<in> A & b \<in> B a" |
|
51 |
by (metis SigmaD1 SigmaD2) |
|
52 |
||
| 36566 | 53 |
lemma "(a, b) \<in> (SIGMA x:A. {y. x = f y}) \<Longrightarrow> a \<in> A \<and> a = f b"
|
54 |
(* Metis says this is satisfiable! |
|
| 29676 | 55 |
by (metis CollectD SigmaD1 SigmaD2) |
56 |
*) |
|
| 23449 | 57 |
by (meson CollectD SigmaD1 SigmaD2) |
58 |
||
| 36566 | 59 |
lemma "(a, b) \<in> (SIGMA x:A. {y. x = f y}) \<Longrightarrow> a \<in> A \<and> a = f b"
|
60 |
by (metis mem_Sigma_iff singleton_conv2 vimage_Collect_eq vimage_singleton_eq) |
|
| 24827 | 61 |
|
| 36566 | 62 |
lemma "(a, b) \<in> (SIGMA x:A. {y. x = f y}) \<Longrightarrow> a \<in> A \<and> a = f b"
|
63 |
proof - |
|
64 |
assume A1: "(a, b) \<in> (SIGMA x:A. {y. x = f y})"
|
|
65 |
hence F1: "a \<in> A" by (metis mem_Sigma_iff) |
|
66 |
have "b \<in> {R. a = f R}" by (metis A1 mem_Sigma_iff)
|
|
67 |
hence F2: "b \<in> (\<lambda>R. a = f R)" by (metis Collect_def) |
|
68 |
hence "a = f b" by (unfold mem_def) |
|
69 |
thus "a \<in> A \<and> a = f b" by (metis F1) |
|
| 24827 | 70 |
qed |
| 23449 | 71 |
|
72 |
lemma "(cl,f) \<in> CLF ==> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) ==> f \<in> pset cl"
|
|
| 24827 | 73 |
by (metis Collect_mem_eq SigmaD2) |
| 23449 | 74 |
|
|
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24632
diff
changeset
|
75 |
lemma "(cl,f) \<in> CLF ==> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) ==> f \<in> pset cl"
|
| 36566 | 76 |
proof - |
77 |
assume A1: "(cl, f) \<in> CLF" |
|
78 |
assume A2: "CLF = (SIGMA cl:CL. {f. f \<in> pset cl})"
|
|
79 |
have F1: "\<forall>v. (\<lambda>R. R \<in> v) = v" by (metis Collect_mem_eq Collect_def) |
|
80 |
have "\<forall>v u. (u, v) \<in> CLF \<longrightarrow> v \<in> {R. R \<in> pset u}" by (metis A2 mem_Sigma_iff)
|
|
81 |
hence "\<forall>v u. (u, v) \<in> CLF \<longrightarrow> v \<in> pset u" by (metis F1 Collect_def) |
|
82 |
hence "f \<in> pset cl" by (metis A1) |
|
83 |
thus "f \<in> pset cl" by metis |
|
| 24827 | 84 |
qed |
| 23449 | 85 |
|
86 |
lemma |
|
| 43197 | 87 |
"(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==>
|
| 23449 | 88 |
f \<in> pset cl \<rightarrow> pset cl" |
| 45562 | 89 |
by (metis (no_types) Collect_def Sigma_triv mem_def) |
90 |
||
91 |
lemma |
|
92 |
"(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==>
|
|
93 |
f \<in> pset cl \<rightarrow> pset cl" |
|
| 36566 | 94 |
proof - |
95 |
assume A1: "(cl, f) \<in> (SIGMA cl:CL. {f. f \<in> pset cl \<rightarrow> pset cl})"
|
|
96 |
have F1: "\<forall>v. (\<lambda>R. R \<in> v) = v" by (metis Collect_mem_eq Collect_def) |
|
97 |
have "f \<in> {R. R \<in> pset cl \<rightarrow> pset cl}" using A1 by simp
|
|
98 |
hence "f \<in> pset cl \<rightarrow> pset cl" by (metis F1 Collect_def) |
|
99 |
thus "f \<in> pset cl \<rightarrow> pset cl" by metis |
|
| 24827 | 100 |
qed |
| 23449 | 101 |
|
102 |
lemma |
|
103 |
"(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
|
|
| 45562 | 104 |
f \<in> pset cl \<inter> cl" |
105 |
by (metis (no_types) Collect_conj_eq Int_def Sigma_triv inf_idem) |
|
106 |
||
107 |
lemma |
|
108 |
"(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
|
|
109 |
f \<in> pset cl \<inter> cl" |
|
| 36566 | 110 |
proof - |
111 |
assume A1: "(cl, f) \<in> (SIGMA cl:CL. {f. f \<in> pset cl \<inter> cl})"
|
|
112 |
have F1: "\<forall>v. (\<lambda>R. R \<in> v) = v" by (metis Collect_mem_eq Collect_def) |
|
113 |
have "f \<in> {R. R \<in> pset cl \<inter> cl}" using A1 by simp
|
|
114 |
hence "f \<in> Id_on cl `` pset cl" by (metis F1 Int_commute Image_Id_on Collect_def) |
|
115 |
hence "f \<in> Id_on cl `` pset cl" by metis |
|
116 |
hence "f \<in> cl \<inter> pset cl" by (metis Image_Id_on) |
|
117 |
thus "f \<in> pset cl \<inter> cl" by (metis Int_commute) |
|
| 24827 | 118 |
qed |
119 |
||
| 23449 | 120 |
lemma |
121 |
"(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) ==>
|
|
122 |
(f \<in> pset cl \<rightarrow> pset cl) & (monotone f (pset cl) (order cl))" |
|
123 |
by auto |
|
124 |
||
| 43197 | 125 |
lemma "(cl,f) \<in> CLF ==> |
| 23449 | 126 |
CLF \<subseteq> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
|
127 |
f \<in> pset cl \<inter> cl" |
|
| 24827 | 128 |
by auto |
| 27368 | 129 |
|
| 43197 | 130 |
lemma "(cl,f) \<in> CLF ==> |
| 23449 | 131 |
CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
|
132 |
f \<in> pset cl \<inter> cl" |
|
| 24827 | 133 |
by auto |
| 36566 | 134 |
|
| 43197 | 135 |
lemma |
136 |
"(cl,f) \<in> CLF ==> |
|
137 |
CLF \<subseteq> (SIGMA cl': CL. {f. f \<in> pset cl' \<rightarrow> pset cl'}) ==>
|
|
| 23449 | 138 |
f \<in> pset cl \<rightarrow> pset cl" |
| 31754 | 139 |
by fast |
| 36566 | 140 |
|
| 43197 | 141 |
lemma |
142 |
"(cl,f) \<in> CLF ==> |
|
143 |
CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==>
|
|
| 23449 | 144 |
f \<in> pset cl \<rightarrow> pset cl" |
| 24827 | 145 |
by auto |
| 36566 | 146 |
|
| 43197 | 147 |
lemma |
148 |
"(cl,f) \<in> CLF ==> |
|
| 23449 | 149 |
CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) ==>
|
150 |
(f \<in> pset cl \<rightarrow> pset cl) & (monotone f (pset cl) (order cl))" |
|
151 |
by auto |
|
152 |
||
153 |
lemma "map (%x. (f x, g x)) xs = zip (map f xs) (map g xs)" |
|
154 |
apply (induct xs) |
|
| 45562 | 155 |
apply (metis map.simps(1) zip_Nil) |
156 |
by (metis (lam_lifting, no_types) map.simps(2) zip_Cons_Cons) |
|
| 23449 | 157 |
|
| 43197 | 158 |
lemma "map (%w. (w -> w, w \<times> w)) xs = |
| 23449 | 159 |
zip (map (%w. w -> w) xs) (map (%w. w \<times> w) xs)" |
160 |
apply (induct xs) |
|
| 45562 | 161 |
apply (metis map.simps(1) zip_Nil) |
| 36566 | 162 |
by auto |
| 23449 | 163 |
|
| 45562 | 164 |
lemma "(%x. Suc (f x)) ` {x. even x} <= A ==> \<forall>x. even x --> Suc (f x) \<in> A"
|
165 |
by (metis Collect_def image_eqI mem_def subsetD) |
|
| 23449 | 166 |
|
| 43197 | 167 |
lemma "(%x. f (f x)) ` ((%x. Suc(f x)) ` {x. even x}) <= A
|
| 45503 | 168 |
==> (\<forall>x. even x --> f (f (Suc(f x))) \<in> A)" |
| 45562 | 169 |
by (metis Collect_def imageI mem_def set_rev_mp) |
| 23449 | 170 |
|
| 43197 | 171 |
lemma "f \<in> (%u v. b \<times> u \<times> v) ` A ==> \<forall>u v. P (b \<times> u \<times> v) ==> P(f y)" |
| 45562 | 172 |
(* sledgehammer *) |
| 23449 | 173 |
by auto |
174 |
||
175 |
lemma image_TimesA: "(%(x,y). (f x, g y)) ` (A \<times> B) = (f`A) \<times> (g`B)" |
|
| 45562 | 176 |
by (metis map_pair_def map_pair_surj_on) |
| 23449 | 177 |
|
178 |
lemma image_TimesB: |
|
| 36566 | 179 |
"(%(x,y,z). (f x, g y, h z)) ` (A \<times> B \<times> C) = (f`A) \<times> (g`B) \<times> (h`C)" |
| 45562 | 180 |
(* sledgehammer *) |
| 23449 | 181 |
by force |
182 |
||
183 |
lemma image_TimesC: |
|
| 43197 | 184 |
"(%(x,y). (x \<rightarrow> x, y \<times> y)) ` (A \<times> B) = |
185 |
((%x. x \<rightarrow> x) ` A) \<times> ((%y. y \<times> y) ` B)" |
|
| 45562 | 186 |
by (metis image_TimesA) |
| 23449 | 187 |
|
188 |
end |