-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathUnion_con_la_imagen.thy
116 lines (103 loc) · 3.01 KB
/
Union_con_la_imagen.thy
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
(* Union_con_la_imagen.thy
-- Unión con la imagen
-- José A. Alonso Jiménez <https://jaalonso.github.io>
-- Sevilla, 22-abril-2024
-- ------------------------------------------------------------------ *)
(* ---------------------------------------------------------------------
-- Demostrar que
-- f ` (s \<union> f -` v) \<subseteq> f ` s \<union> v
-- ------------------------------------------------------------------ *)
theory Union_con_la_imagen
imports Main
begin
(* 1\<ordfeminine> demostración *)
lemma "f ` (s \<union> f -` v) \<subseteq> f ` s \<union> v"
proof (rule subsetI)
fix y
assume "y \<in> f ` (s \<union> f -` v)"
then show "y \<in> f ` s \<union> v"
proof (rule imageE)
fix x
assume "y = f x"
assume "x \<in> s \<union> f -` v"
then show "y \<in> f ` s \<union> v"
proof (rule UnE)
assume "x \<in> s"
then have "f x \<in> f ` s"
by (rule imageI)
with \<open>y = f x\<close> have "y \<in> f ` s"
by (rule ssubst)
then show "y \<in> f ` s \<union> v"
by (rule UnI1)
next
assume "x \<in> f -` v"
then have "f x \<in> v"
by (rule vimageD)
with \<open>y = f x\<close> have "y \<in> v"
by (rule ssubst)
then show "y \<in> f ` s \<union> v"
by (rule UnI2)
qed
qed
qed
(* 2\<ordfeminine> demostración *)
lemma "f ` (s \<union> f -` v) \<subseteq> f ` s \<union> v"
proof
fix y
assume "y \<in> f ` (s \<union> f -` v)"
then show "y \<in> f ` s \<union> v"
proof
fix x
assume "y = f x"
assume "x \<in> s \<union> f -` v"
then show "y \<in> f ` s \<union> v"
proof
assume "x \<in> s"
then have "f x \<in> f ` s" by simp
with \<open>y = f x\<close> have "y \<in> f ` s" by simp
then show "y \<in> f ` s \<union> v" by simp
next
assume "x \<in> f -` v"
then have "f x \<in> v" by simp
with \<open>y = f x\<close> have "y \<in> v" by simp
then show "y \<in> f ` s \<union> v" by simp
qed
qed
qed
(* 3\<ordfeminine> demostración *)
lemma "f ` (s \<union> f -` v) \<subseteq> f ` s \<union> v"
proof
fix y
assume "y \<in> f ` (s \<union> f -` v)"
then show "y \<in> f ` s \<union> v"
proof
fix x
assume "y = f x"
assume "x \<in> s \<union> f -` v"
then show "y \<in> f ` s \<union> v"
proof
assume "x \<in> s"
then show "y \<in> f ` s \<union> v" by (simp add: \<open>y = f x\<close>)
next
assume "x \<in> f -` v"
then show "y \<in> f ` s \<union> v" by (simp add: \<open>y = f x\<close>)
qed
qed
qed
(* 4\<ordfeminine> demostración *)
lemma "f ` (s \<union> f -` v) \<subseteq> f ` s \<union> v"
proof
fix y
assume "y \<in> f ` (s \<union> f -` v)"
then show "y \<in> f ` s \<union> v"
proof
fix x
assume "y = f x"
assume "x \<in> s \<union> f -` v"
then show "y \<in> f ` s \<union> v" using \<open>y = f x\<close> by blast
qed
qed
(* 5\<ordfeminine> demostración *)
lemma "f ` (s \<union> f -` u) \<subseteq> f ` s \<union> u"
by auto
end