-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathMonotonia_de_la_imagen_de_conjuntos.thy
72 lines (65 loc) · 1.73 KB
/
Monotonia_de_la_imagen_de_conjuntos.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
(* Monotonia_de_la_imagen_de_conjuntos.thy
Si s \<subseteq> t, entonces f[s] \<subseteq> f[t].
José A. Alonso Jiménez <https://jaalonso.github.io>
Sevilla, 3-abril-2024
------------------------------------------------------------------ *)
(* ---------------------------------------------------------------------
Demostrar que si s \<subseteq> t, entonces
f ` s \<subseteq> f ` t
------------------------------------------------------------------- *)
theory Monotonia_de_la_imagen_de_conjuntos
imports Main
begin
(* 1\<ordfeminine> demostración *)
lemma
assumes "s \<subseteq> t"
shows "f ` s \<subseteq> f ` t"
proof (rule subsetI)
fix y
assume "y \<in> f ` s"
then show "y \<in> f ` t"
proof (rule imageE)
fix x
assume "y = f x"
assume "x \<in> s"
then have "x \<in> t"
using \<open>s \<subseteq> t\<close> by (simp only: set_rev_mp)
then have "f x \<in> f ` t"
by (rule imageI)
with \<open>y = f x\<close> show "y \<in> f ` t"
by (rule ssubst)
qed
qed
(* 2\<ordfeminine> demostración *)
lemma
assumes "s \<subseteq> t"
shows "f ` s \<subseteq> f ` t"
proof
fix y
assume "y \<in> f ` s"
then show "y \<in> f ` t"
proof
fix x
assume "y = f x"
assume "x \<in> s"
then have "x \<in> t"
using \<open>s \<subseteq> t\<close> by (simp only: set_rev_mp)
then have "f x \<in> f ` t"
by simp
with \<open>y = f x\<close> show "y \<in> f ` t"
by simp
qed
qed
(* 3\<ordfeminine> demostración *)
lemma
assumes "s \<subseteq> t"
shows "f ` s \<subseteq> f ` t"
using assms
by blast
(* 4\<ordfeminine> demostración *)
lemma
assumes "s \<subseteq> t"
shows "f ` s \<subseteq> f ` t"
using assms
by (simp only: image_mono)
end