Acciones

Diferencia entre revisiones de «Sol 6»

De Lógica matemática y fundamentos (2018-19)

 
m
(Sin diferencias)

Revisión del 12:51 9 abr 2019

chapter {* R6: Deducción natural en lógica de de primer orden *}

theory R6_sol
imports Main 
begin

text {*
  Demostrar o refutar los siguientes lemas usando sólo las reglas
  básicas de deducción natural de la lógica proposicional, de los
  cuantificadores y de la igualdad: 
  · conjI:      ⟦P; Q⟧ ⟹ P ∧ Q
  · conjunct1:  P ∧ Q ⟹ P
  · conjunct2:  P ∧ Q ⟹ Q  
  · notnotD:    ¬¬ P ⟹ P
  · mp:         ⟦P ⟶ Q; P⟧ ⟹ Q 
  · impI:       (P ⟹ Q) ⟹ P ⟶ Q
  · disjI1:     P ⟹ P ∨ Q
  · disjI2:     Q ⟹ P ∨ Q
  · disjE:      ⟦P ∨ Q; P ⟹ R; Q ⟹ R⟧ ⟹ R 
  · FalseE:     False ⟹ P
  · notE:       ⟦¬P; P⟧ ⟹ R
  · notI:       (P ⟹ False) ⟹ ¬P
  · iffI:       ⟦P ⟹ Q; Q ⟹ P⟧ ⟹ P = Q
  · iffD1:      ⟦Q = P; Q⟧ ⟹ P 
  · iffD2:      ⟦P = Q; Q⟧ ⟹ P
  · ccontr:     (¬P ⟹ False) ⟹ P
  . excluded_middel:(¬P ∨ P) 

  · allE:       ⟦∀x. P x; P x ⟹ R⟧ ⟹ R
  · allI:       (⋀x. P x) ⟹ ∀x. P x
  · exI:        P x ⟹ ∃x. P x
  · exE:        ⟦∃x. P x; ⋀x. P x ⟹ Q⟧ ⟹ Q

  · refl:       t = t
  · subst:      ⟦s = t; P s⟧ ⟹ P t
  · trans:      ⟦r = s; s = t⟧ ⟹ r = t
  · sym:        s = t ⟹ t = s
  · not_sym:    t ≠ s ⟹ s ≠ t
  · ssubst:     ⟦t = s; P s⟧ ⟹ P t
  · box_equals: ⟦a = b; a = c; b = d⟧ ⟹ a: = d
  · arg_cong:   x = y ⟹ f x = f y
  · fun_cong:   f = g ⟹ f x = g x
  · cong:       ⟦f = g; x = y⟧ ⟹ f x = g y
*}

text {*
  Se usarán las reglas notnotI y mt que demostramos a continuación.
  *}

lemma notnotI: "P ⟹ ¬¬ P"
by auto

lemma mt: "⟦F ⟶ G; ¬G⟧ ⟹ ¬F"
by auto

text {* --------------------------------------------------------------- 
  Ejercicio 1. Demostrar
       ∀x. P x ⟶ Q x ⊢ (∀x. P x) ⟶ (∀x. Q x)
  ------------------------------------------------------------------ *}

lemma ejercicio_1a: 
  assumes "∀x. P x ⟶ Q x"
  shows   "(∀x. P x) ⟶ (∀x. Q x)"
proof
  assume "∀x. P x"
  show "∀x. Q x"
  proof
    fix a
    have "P a" using `∀x. P x` by (rule allE)
    have "P a ⟶ Q a" using assms ..
    thus "Q a"  using `P a` ..
  qed
qed

lemma ejercicio_1b: 
  assumes "∀x. P x ⟶ Q x"
  shows   "(∀x. P x) ⟶ (∀x. Q x)"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 2. Demostrar
       ∃x. ¬(P x) ⊢ ¬(∀x. P x)
  ------------------------------------------------------------------ *}

lemma ejercicio_2a: 
  assumes "∃x. ¬(P x)"
  shows   "¬(∀x. P x)"
  proof
    assume "∀x. P x"
    obtain a where "¬(P a)" using assms ..
    have "P a" using `∀x. P x` ..
    with `¬(P a)` show False ..
  qed

lemma ejercicio_2b: 
  assumes "∃x. ¬(P x)"
  shows   "¬(∀x. P x)"
  using assms by auto


text {* --------------------------------------------------------------- 
  Ejercicio 3. Demostrar
       ∀x. P x ⊢ ∀y. P y
  ------------------------------------------------------------------ *}

lemma ejercicio_3a: 
  assumes "∀x. P x"
  shows   "∀y. P y"
proof
  fix a show "P a" using assms ..
qed

lemma ejercicio_3b: 
  assumes "∀x. P x"
  shows   "∀y. P y"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 4. Demostrar
       ∀x. P x ⟶ Q x ⊢ (∀x. ¬(Q x)) ⟶ (∀x. ¬ (P x))
  ------------------------------------------------------------------ *}

lemma ejercicio_4a: 
  assumes "∀x. P x ⟶ Q x"
  shows   "(∀x. ¬(Q x)) ⟶ (∀x. ¬ (P x))"
proof
  assume "∀x. ¬(Q x)"
  show "∀x. ¬ (P x)"
  proof
    fix a
    have "¬(Q a)" using `∀x. ¬(Q x)` ..
    have "P a ⟶ Q a" using assms ..
    thus "¬(P a)" using `¬(Q a)`  by (rule mt)
  qed
qed

lemma ejercicio_4b: 
  assumes "∀x. P x ⟶ Q x"
  shows   "(∀x. ¬(Q x)) ⟶ (∀x. ¬ (P x))"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 5. Demostrar
       ∀x. P x  ⟶ ¬(Q x) ⊢ ¬(∃x. P x ∧ Q x)
  ------------------------------------------------------------------ *}

lemma ejercicio_5a: 
  assumes "∀x. P x  ⟶ ¬(Q x)"
  shows   "¬(∃x. P x ∧ Q x)"
  proof
    assume "∃x. P x ∧ Q x"
    then obtain a where "P a ∧ Q a" ..
    hence "P a" ..
    have "P a  ⟶ ¬(Q a)" using assms ..
    hence "¬(Q a)" using `P a` ..
    have "Q a" using `P a ∧ Q a` ..
    with `¬(Q a)` show False ..
  qed

lemma ejercicio_5b: 
  assumes "∀x. P x  ⟶ ¬(Q x)"
  shows   "¬(∃x. P x ∧ Q x)"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 6. Demostrar
       ∀x y. P x y ⊢ ∀u v. P u v
  ------------------------------------------------------------------ *}

lemma ejercicio_6a: 
  assumes "∀x y. P x y"
  shows   "∀u v. P u v"
proof (rule allI)+
  fix a b
  have "∀y. P a y" using assms ..
  thus "P a b" ..
qed


lemma ejercicio_6b: 
  assumes "∀x y. P x y"
  shows   "∀u v. P u v"
proof (rule allI)
  fix a
  show "∀y. P a y" using assms ..
qed


lemma ejercicio_6c: 
  assumes "∀x y. P x y"
  shows   "∀u v. P u v"
proof
  fix a
  show "∀y. P a y" using assms ..
qed

lemma ejercicio_6d: 
  assumes "∀x y. P x y"
  shows   "∀u v. P u v"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 7. Demostrar
       ∃x y. P x y ⟹ ∃u v. P u v
  ------------------------------------------------------------------ *}

lemma ejercicio_7a: 
  assumes "∃x y. P x y"
  shows   "∃u v. P u v"
  proof -
    obtain a where "∃y. P a y" using assms ..
    then obtain b where "P a b" ..
    hence "∃v. P a v" ..
    thus "∃u v. P u v" ..
  qed

lemma ejercicio_7c: 
  assumes "∃x y. P x y"
  shows   "∃u v. P u v"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 8. Demostrar
       ∃x. ∀y. P x y ⊢ ∀y. ∃x. P x y
  ------------------------------------------------------------------ *}

lemma ejercicio_8a: 
  assumes "∃x. ∀y. P x y"
  shows   "∀y. ∃x. P x y"
proof
  fix b
  obtain a where "∀y. P a y" using assms ..
  hence "P a b" ..
  thus "∃x. P x b" ..
qed


lemma ejercicio_8b: 
  assumes "∃x. ∀y. P x y"
  shows   "∀y. ∃x. P x y"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 9. Demostrar
       ∃x. P a ⟶ Q x ⊢ P a ⟶ (∃x. Q x)
  ------------------------------------------------------------------ *}

lemma ejercicio_9a: 
  assumes "∃x. P a ⟶ Q x"
  shows   "P a ⟶ (∃x. Q x)"
  proof
    assume "P a"
    obtain b where "P a ⟶ Q b" using assms ..
    hence "Q b" using `P a`..
    thus "∃x . Q x" ..
  qed

lemma ejercicio_9b: 
  assumes "∃x. P a ⟶ Q x"
  shows   "P a ⟶ (∃x. Q x)"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 10. Demostrar
       P a ⟶ (∃x. Q x) ⊢ ∃x. P a ⟶ Q x 
  ------------------------------------------------------------------ *}

lemma ejercicio_10a: 
  fixes P Q :: "'b ⇒ bool"
  assumes "P a ⟶ (∃x. Q x)"
  shows   "∃x. P a ⟶ Q x"
  proof -
    have "¬(P a) ∨ (P a)" ..
    thus "∃x. P a ⟶ Q x"
      proof
        assume "¬(P a)"
        have "P a ⟶ Q a"
          proof
            assume "P a"
            with `¬(P a)` show "Q a" ..
          qed
          thus "∃x. P a ⟶ Q x" ..
      next
        assume "P a"
        with assms have "∃x. Q x" ..
        then obtain b where "Q b" ..
        hence "P a ⟶  Q b" ..
        thus "∃x. P a ⟶ Q x" ..
      qed
  qed


lemma ejercicio_10b: 
  fixes P Q :: "'b ⇒ bool" 
  assumes "P a ⟶ (∃x. Q x)"
  shows   "∃x. P a ⟶ Q x"
  using assms by auto


text {* --------------------------------------------------------------- 
  Ejercicio 11. Demostrar
       (∃x. P x) ⟶ Q a ⊢ ∀x. P x ⟶ Q a
  ------------------------------------------------------------------ *}

lemma ejercicio_11a: 
  assumes "(∃x. P x) ⟶ Q a"
  shows   "∀x. P x ⟶ Q a"
proof
  fix b
  show "P b ⟶  Q a"
    proof
      assume "P b"
      hence "∃x. P x" ..
      with assms show "Q a" ..
    qed
qed


lemma ejercicio_11b: 
  assumes "(∃x. P x) ⟶ Q a"
  shows   "∀x. P x ⟶ Q a"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 12. Demostrar
       ∀x. P x ⟶ Q a ⊢ ∃ x. P x ⟶ Q a
  ------------------------------------------------------------------ *}

lemma ejercicio_12a: 
  assumes "∀x. P x ⟶ Q a"
  shows   "∃x. P x ⟶ Q a"
proof - (* comentar sin - *)
  have "P b ⟶  Q a" using assms .. (* ¿universo? *)
  thus "∃x. P x ⟶ Q a" ..
qed


lemma ejercicio_12b: 
  assumes "∀x. P x ⟶ Q a"
  shows   "∃x. P x ⟶ Q a"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 13. Demostrar
       (∀x. P x) ∨ (∀x. Q x) ⊢ ∀x. P x ∨ Q x
  ------------------------------------------------------------------ *}

lemma ejercicio_13a: 
  assumes "(∀x. P x) ∨ (∀x. Q x)"
  shows   "∀x. P x ∨ Q x"
  using assms
proof
  assume "∀x. P x"
  show "∀x. P x ∨ Q x"
    proof
      fix a
      have "P a" using `∀x. P x` ..
      thus "P a ∨ Q a" ..
    qed
next
  assume "∀x. Q x"
  show "∀x. P x ∨ Q x"
    proof
      fix a
      have "Q a" using `∀x. Q x` ..
      thus "P a ∨ Q a" ..
    qed
qed

lemma ejercicio_13b: 
  assumes "(∀x. P x) ∨ (∀x. Q x)"
  shows   "∀x. P x ∨ Q x"
proof
  fix a
  note assms
  thus "P a ∨ Q a"
  proof
    assume "∀x. P x"
    hence "P a" ..
    thus "P a ∨ Q a" ..
  next
    assume "∀x. Q x"
    hence "Q a" ..
    thus "P a ∨ Q a" ..
  qed
qed

lemma ejercicio_13c: 
  assumes "(∀x. P x) ∨ (∀x. Q x)"
  shows   "∀x. P x ∨ Q x"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 14. Demostrar
       ∃x. P x ∧ Q x ⊢ (∃x. P x) ∧ (∃x. Q x)
  ------------------------------------------------------------------ *}

lemma ejercicio_14a: 
  assumes "∃x. P x ∧ Q x"
  shows   "(∃x. P x) ∧ (∃x. Q x)"
proof
  show "∃x. P x"
    proof -
  obtain a where "P a ∧ Q a" using assms ..
  hence "P a" ..
  thus "∃x. P x" ..
    qed
next
  show "∃x. Q x"
    proof -
      obtain a where "P a ∧ Q a" using assms ..
      hence "Q a" ..
      thus "∃x. Q x" ..
    qed
qed

lemma ejercicio_14b: 
  assumes "∃x. P x ∧ Q x"
  shows   "(∃x. P x) ∧ (∃x. Q x)"
proof
  obtain a where "P a ∧ Q a" using assms ..
  hence "P a" ..
  thus "∃x. P x" ..
next
  obtain a where "P a ∧ Q a" using assms ..
  hence "Q a" ..
  thus "∃x. Q x" ..
qed



lemma ejercicio_14c:
  assumes "∃x. P x ∧ Q x"
  shows   "(∃x. P x) ∧ (∃x. Q x)"
proof -
  obtain a where "P a ∧ Q a" using assms ..
  {have "∃x. P x"
    proof -
      have "P a" using `P a ∧ Q a` ..
      thus "∃x. P x" ..
    qed}
moreover
  {have "∃x. Q x"
  proof -
    have "Q a" using `P a ∧ Q a` ..
    thus "∃x. Q x" ..
  qed}
  ultimately show "(∃x. P x) ∧ (∃x. Q x)" ..
qed


lemma ejercicio_14d:
  assumes "∃x. P x ∧ Q x"
  shows   "(∃x. P x) ∧ (∃x. Q x)"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 15. Demostrar
       ∀x y. P y ⟶ Q x ⊢ (∃y. P y) ⟶ (∀x. Q x)
  ------------------------------------------------------------------ *}

lemma ejercicio_15a: 
  assumes "∀x y. P y ⟶ Q x"
  shows   "(∃y. P y) ⟶ (∀x. Q x)"
proof
  assume "∃y. P y"
  then obtain b where "P b" ..
  show "∀x. Q x"
  proof
    fix a
    have "∀y. P y ⟶ Q a" using assms ..
    hence "P b ⟶ Q a" ..
    thus "Q a" using `P b` ..
  qed
qed

lemma ejercicio_15b: 
  assumes "∀x y. P y ⟶ Q x"
  shows   "(∃y. P y) ⟶ (∀x. Q x)"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 16. Demostrar
       ¬(∀x. ¬(P x)) ⊢ ∃x. P x
  ------------------------------------------------------------------ *}

lemma ejercicio_16a: 
  assumes "¬(∀x. ¬(P x))"
  shows   "∃x. P x"
proof (rule ccontr)
  assume "¬(∃x. P x)"
  have "∀x. ¬(P x)"
    proof
      fix a
      show "¬(P a)"
        proof
          assume "P a"
          hence "∃x. P x" ..
          with `¬(∃x. P x)` show False ..
        qed
    qed
    with assms show False ..
qed


lemma ejercicio_16b: 
  assumes "¬(∀x. ¬(P x))"
  shows   "∃x. P x"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 17. Demostrar
       ∀x. ¬(P x) ⊢ ¬(∃x. P x)
  ------------------------------------------------------------------ *}

lemma ejercicio_17a: 
  assumes "∀x. ¬(P x)"
  shows   "¬(∃x. P x)"
proof
  assume "∃x. P x"
  then obtain a where "P a " ..
  have "¬(P a)" using assms ..
  thus False using `P a` ..
qed

lemma ejercicio_17b: 
  assumes "∀x. ¬(P x)"
  shows   "¬(∃x. P x)"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 18. Demostrar
       ∃x. P x ⊢ ¬(∀x. ¬(P x))
  ------------------------------------------------------------------ *}

lemma ejercicio_18a: 
  assumes "∃x. P x"
  shows   "¬(∀x. ¬(P x))"
proof
  obtain a where "P a" using assms ..
  assume "∀x. ¬(P x)"
  hence "¬(P a)" ..
  thus False using `P a` ..
qed

lemma ejercicio_18b: 
  assumes "∃x. P x"
  shows   "¬(∀x. ¬(P x))"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 19. Demostrar
       P a ⟶ (∀x. Q x) ⊢ ∀x. P a ⟶ Q x
  ------------------------------------------------------------------ *}

lemma ejercicio_19a: 
  assumes "P a ⟶ (∀x. Q x)"
  shows   "∀x. P a ⟶ Q x"
proof
  fix b
  show "P a ⟶ Q b"
  proof
    assume "P a"
    with assms have "∀x. Q x" ..
    thus "Q b" ..
  qed 
qed

lemma ejercicio_19b: 
  assumes "P a ⟶ (∀x. Q x)"
  shows   "∀x. P a ⟶ Q x"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 20. Demostrar
       {∀x y z. R x y ∧ R y z ⟶ R x z, 
        ∀x. ¬(R x x)}
       ⊢ ∀x y. R x y ⟶ ¬(R y x)
  ------------------------------------------------------------------ *}

lemma ejercicio_20a: 
  assumes "∀x y z. R x y ∧ R y z ⟶ R x z"
          "∀x. ¬(R x x)" 
  shows   "∀x y. R x y ⟶ ¬(R y x)"
proof (rule allI)+
  fix a b
  show "R a b ⟶ ¬(R b a)"
  proof
    assume "R a b"
    show "¬(R b a)"
      proof
        assume "R b a"
        with `R a b` have "R a b ∧ R b a" ..
        have "¬(R a a)" using assms(2) ..
        have "∀y z. R a y ∧ R y z ⟶ R a z" using assms(1) ..
        hence "∀z. R a b ∧ R b z ⟶ R a z"  ..
        hence "R a b ∧ R b a ⟶ R a a" ..
        hence "R a a" using `R a b ∧ R b a` ..
        with `¬(R a a)` show False ..
      qed
  qed
qed


lemma ejercicio_20b:
  "⟦∀x y z. R x y ∧ R y z ⟶ R x z; ∀x. ¬(R x x)⟧ ⟹ ∀x y. R x y ⟶ ¬(R y x)"
by metis

text {* --------------------------------------------------------------- 
  Ejercicio 21. Demostrar
     {∀x. P x ∨ Q x, ∃x. ¬(Q x), ∀x. R x ⟶ ¬(P x)} ⊢ ∃x. ¬(R x)
  ------------------------------------------------------------------ *}

lemma ejercicio_21a:
  assumes "∀x. P x ∨ Q x" 
          "∃x. ¬(Q x)" 
          "∀x. R x ⟶ ¬(P x)"
  shows   "∃x. ¬(R x)" 
proof -
  obtain a where "¬(Q a)" using assms(2) ..
  have "P a ∨ Q a" using assms(1) ..
  hence "¬(R a)" 
    proof 
      assume "P a"
      hence "¬¬(P a)" by (rule notnotI)
      have "R a ⟶ ¬(P a)" using assms(3) ..
      thus "¬(R a)" using `¬¬(P a)` by (rule mt)
    next
      assume "Q a"
      with `¬(Q a)` show "¬(R a)" ..
    qed
    thus "∃x. ¬(R x)"  ..
qed

lemma ejercicio_21b:
  assumes "∀x. P x ∨ Q x" 
          "∃x. ¬(Q x)" 
          "∀x. R x ⟶ ¬(P x)"
  shows   "∃x. ¬(R x)" 
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 22. Demostrar
     {∀x. P x ⟶ Q x ∨ R x, ¬(∃x. P x ∧ R x)} ⊢ ∀x. P x ⟶ Q x
  ------------------------------------------------------------------ *}

lemma ejercicio_22a:
  assumes "∀x. P x ⟶ Q x ∨ R x" 
          "¬(∃x. P x ∧ R x)"
  shows   "∀x. P x ⟶ Q x"
proof
  fix a
  show "P a ⟶ Q a"
    proof
      assume "P a"
      have "P a ⟶ Q a ∨ R a" using assms(1) ..
      hence "Q a ∨ R a" using `P a` ..
      thus "Q a"
        proof
          assume "Q a" thus "Q a" .
        next
          assume "R a"
          with `P a` have "P a ∧ R a" ..
          hence "∃x. P x ∧ R x" ..
          with assms(2) show "Q a" ..
        qed
    qed
qed

lemma ejercicio_22b:
  assumes "∀x. P x ⟶ Q x ∨ R x" 
          "¬(∃x. P x ∧ R x)"
  shows   "∀x. P x ⟶ Q x"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 23. Demostrar
     ∃x y. R x y ∨ R y x ⊢ ∃x y. R x y
  ------------------------------------------------------------------ *}

lemma ejercicio_23a:
  assumes "∃x y. R x y ∨ R y x"
  shows   "∃x y. R x y"
proof -
  obtain a where "∃y. R a y ∨ R y a" using assms ..
  then obtain b where "R a b ∨ R b a" ..
  thus "∃x y. R x y"
  proof
    assume "R a b"
    hence "∃y. R a y" ..
    thus "∃x y. R x y" ..
  next
    assume "R b a"
    hence "∃y. R b y" ..
    thus "∃x y. R x y" ..
  qed
qed


lemma ejercicio_23b:
  assumes "∃x y. R x y ∨ R y x"
  shows   "∃x y. R x y"
  using assms by auto

text {* --------------------------------------------------------------- 
  Ejercicio 24. Demostrar
       (∃x. ∀y. P x y) ⟶ (∀y. ∃x. P x y)
  ------------------------------------------------------------------ *}

lemma ejercicio_24a: 
  "(∃x. ∀y. P x y) ⟶ (∀y. ∃x. P x y)"
proof
  assume "∃x. ∀y. P x y"
  then obtain a where "∀y. P a y" ..
  show "∀y. ∃x. P x y"
    proof
      fix b
      have "P a b" using `∀y. P a y` ..
      thus "∃x. P x b" ..
    qed
qed


lemma ejercicio_24b: 
  "(∃x. ∀y. P x y) ⟶ (∀y. ∃x. P x y)"
  by auto

text {* --------------------------------------------------------------- 
  Ejercicio 25. Demostrar
       (∀x. P x ⟶ Q) ⟷ ((∃x. P x) ⟶ Q)
  ------------------------------------------------------------------ *}


lemma ejercicio_25a: 
  "(∀x. P x ⟶ Q) ⟷ ((∃x. P x) ⟶ Q)"
proof
  assume "∀x. P x ⟶ Q"
  show "(∃x. P x) ⟶ Q"
    proof
      assume "∃x. P x"
      then obtain a where "P a" ..
      have "P a ⟶ Q" using `∀x. P x ⟶ Q` ..
      thus "Q" using `P a` ..
    qed
next
  assume "(∃x. P x) ⟶ Q"
  show "∀x. P x ⟶ Q"
    proof
      fix a
      show "P a ⟶ Q"
        proof
          assume "P a"
          hence "∃x. P x" ..
          with `(∃x. P x) ⟶ Q` show "Q" ..
        qed
    qed
qed

lemma ejercicio_25b: 
  "(∀x. P x ⟶ Q) ⟷ ((∃x. P x) ⟶ Q)"
  by auto

text {* --------------------------------------------------------------- 
  Ejercicio 26. Demostrar
       ((∀x. P x) ∧ (∀x. Q x)) ⟷ (∀x. P x ∧ Q x)
  ------------------------------------------------------------------ *}

lemma ejercicio_26a: 
  "((∀x. P x) ∧ (∀x. Q x)) ⟷ (∀x. P x ∧ Q x)"
proof
  assume "(∀x. P x) ∧ (∀x. Q x)"
  hence "∀x. P x" ..
  have "∀x. Q x" using `(∀x. P x) ∧ (∀x. Q x)` ..
  show "∀x. P x ∧ Q x"
  proof
    fix a
    have "Q a" using `∀x. Q x` ..
    have "P a" using `∀x. P x` ..
    thus "P a ∧ Q a"  using `Q a` ..
  qed
  next
    assume "∀x. P x ∧ Q x"
    show "(∀x. P x) ∧ (∀x. Q x)"
    proof
      show "∀x. P x"
      proof
        fix a
        have "P a ∧ Q a" using `∀x. P x ∧ Q x` ..
        thus "P a" ..
      qed
    next
      show "∀x. Q x"
      proof
        fix a
        have "P a ∧ Q a" using `∀x. P x ∧ Q x` ..
        thus "Q a" ..
      qed
    qed
qed

lemma ejercicio_26b: 
  "((∀x. P x) ∧ (∀x. Q x)) ⟷ (∀x. P x ∧ Q x)"
  by auto

text {* --------------------------------------------------------------- 
  Ejercicio 27. Demostrar o refutar
       ((∀x. P x) ∨ (∀x. Q x)) ⟷ (∀x. P x ∨ Q x)
  ------------------------------------------------------------------ *}

lemma ejercicio_27a: 
  "((∀x. P x) ∨ (∀x. Q x)) ⟷ (∀x. P x ∨ Q x)"
  nitpick


(* Nitpicking formula... *)

(* Nitpick found a counterexample for card 'a = 2: *)

(*   Free variables: *)
(*     P = (λx. _)(a\<^isub>1 := False, a\<^isub>2 := True) *)
(*     Q = (λx. _)(a\<^isub>1 := True, a\<^isub>2 := False) *)
oops


text {* --------------------------------------------------------------- 
  Ejercicio 28. Demostrar o refutar
       ((∃x. P x) ∨ (∃x. Q x)) ⟷ (∃x. P x ∨ Q x)
  ------------------------------------------------------------------ *}

lemma ejercicio_28a: 
  "((∃x. P x) ∨ (∃x. Q x)) ⟷ (∃x. P x ∨ Q x)"
by auto

lemma ejercicio_28b: 
  "((∃x. P x) ∨ (∃x. Q x)) ⟷ (∃x. P x ∨ Q x)"
proof
  assume "(∃x. P x) ∨ (∃x. Q x)"
  thus "∃x. P x ∨ Q x"
    proof
      assume "∃x. P x"
      then obtain a where "P a" ..
      hence "P a ∨ Q a" ..
      thus "∃x. P x ∨ Q x" ..
    next
      assume "∃x. Q x"
      then obtain a where "Q a" ..
      hence "P a ∨ Q a" ..
      thus "∃x. P x ∨ Q x" ..
    qed
  next
    assume "∃x. P x ∨ Q x"
    then obtain a where "P a ∨ Q a" ..
    thus "(∃x. P x) ∨ (∃x. Q x)"
      proof
        assume "P a"
        hence "∃x. P x" ..
        thus "(∃x. P x) ∨ (∃x. Q x)" ..
      next
        assume "Q a"
        hence "∃x. Q x" ..
        thus "(∃x. P x) ∨ (∃x. Q x)" ..
      qed
  qed

text {* --------------------------------------------------------------- 
  Ejercicio 29. Demostrar o refutar
       (∀x. ∃y. P x y) ⟶ (∃y. ∀x. P x y)
  ------------------------------------------------------------------ *}

lemma ejercicio_29: 

  "(∀x. ∃y. P x y) ⟶ (∃y. ∀x. P x y)"
nitpick

(*
Nitpicking formula...

Nitpick found a counterexample for card 'a = 2 and card 'b = 2:

  Free variable:
    P = (λx. _)
        (a\<^isub>1 := (λx. _)(b\<^isub>1 := False, b\<^isub>2 := True),
         a\<^isub>2 := (λx. _)(b\<^isub>1 := True, b\<^isub>2 := False))
  Skolem constants:
    λy. x = (λx. _)(b\<^isub>1 := a\<^isub>1, b\<^isub>2 := a\<^isub>2)
    λx. y = (λx. _)(a\<^isub>1 := b\<^isub>2, a\<^isub>2 := b\<^isub>1)
*)
oops

text {* --------------------------------------------------------------- 
  Ejercicio 30. Demostrar o refutar
       (¬(∀x. P x)) ⟷ (∃x. ¬P x)
  ------------------------------------------------------------------ *}

lemma ejercicio_30a: 
  "(¬(∀x. P x)) ⟷ (∃x. ¬P x)"
  by auto

lemma ejercicio_30b: 
  "(¬(∀x. P x)) ⟷ (∃x. ¬P x)"
proof
  assume "¬(∀x. P x)"
  show "∃x. ¬P x"
  proof (rule ccontr)
    assume "¬(∃x. ¬P x)"
    have "∀x. P x"
      proof
        fix a show "P a"
          proof (rule ccontr)
            assume "¬(P a)"
            hence "∃x. ¬P x" ..
            with `¬(∃x. ¬P x)` show False ..
          qed
      qed
      with `¬(∀x. P x)` show False .. 
  qed
next
  assume "∃x. ¬P x"
  then obtain a where "¬ (P a)" ..
  show  "¬ (∀x. P x)"
    proof
      assume "∀x. P x"
      hence "P a" ..
      with `¬ (P a)` show False ..
    qed
qed

section {* Ejercicios sobre igualdad *}

text {* --------------------------------------------------------------- 
  Ejercicio 31. Demostrar o refutar
       P a ⟹ ∀x. x = a ⟶ P x
  ------------------------------------------------------------------ *}

lemma ejercicio_31a:
  assumes "P a"
  shows   "∀x. x = a ⟶ P x"
  using assms by auto

lemma ejercicio_31b:
  assumes "P a"
  shows   "∀x. x = a ⟶ P x"
proof
  fix b 
  show "b = a ⟶ P b"
  proof
  assume "b = a" thus "P b" using assms by (rule ssubst)
  qed
qed

text {* --------------------------------------------------------------- 
  Ejercicio 32. Demostrar o refutar
       ∃x y. R x y ∨ R y x; ¬(∃x. R x x)⟧ ⟹ ∃x y. x ≠ y
  ------------------------------------------------------------------ *}

lemma ejercicio_32a:
  fixes R :: "'c ⇒ 'c ⇒ bool"
  assumes "∃x y. R x y ∨ R y x"
          "¬(∃x. R x x)"
  shows   "∃(x::'c) y. x ≠ y"
  using assms by metis

lemma ejercicio_32b:
  fixes R :: "'c ⇒ 'c ⇒ bool"
  assumes "∃x y. R x y ∨ R y x"
          "¬(∃x. R x x)"
  shows   "∃(x::'c) y. x ≠ y"
proof -
  from assms(1) obtain a where "∃y. R a y ∨ R y a" ..
  then obtain b where "R a b ∨ R b a" ..
  thus "∃(x::'c) y. x ≠ y"
  proof
    assume "R a b"
    have "¬(a = b)"
      proof 
        assume "a = b"
        hence "R a a" using `R a b` by (rule ssubst)
        hence "∃x. R x x" ..
        with assms(2) show False ..
      qed
      hence "∃y. a ≠ y" ..
      thus "∃(x::'c) y. x ≠ y" ..
  next
    assume "R b a"
    have "¬(a = b)"
      proof 
        assume "a = b"
        hence "R a a" using `R b a` by (rule ssubst)
        hence "∃x. R x x" ..
        with assms(2) show False ..
      qed
      hence "∃y. a ≠ y" ..
      thus "∃(x::'c) y. x ≠ y" ..
  qed
qed

text {* --------------------------------------------------------------- 
  Ejercicio 33. Demostrar o refutar
     {∀x. P a x x, 
      ∀x y z. P x y z ⟶ P (f x) y (f z)} 
     ⊢ P (f a) a (f a)
  ------------------------------------------------------------------ *}

lemma ejercicio_33a:
  assumes "∀x. P a x x"
          "∀x y z. P x y z ⟶ P (f x) y (f z)"
  shows   "P (f a) a (f a)"
  using assms by auto

lemma ejercicio_33b:
  assumes "∀x. P a x x"
          "∀x y z. P x y z ⟶ P (f x) y (f z)"
  shows   "P (f a) a (f a)"
proof -
  have "P a a a" using assms(1) ..
  have "∀y z. P a y z ⟶ P (f a) y (f z)" using assms(2) ..
  hence "∀z. P a a z ⟶ P (f a) a (f z)" ..
  hence "P a a a ⟶ P (f a) a (f a)" ..
  thus "P (f a) a (f a)" using `P a a a` ..
qed

text {* --------------------------------------------------------------- 
  Ejercicio 34. Demostrar o refutar
     {∀x. P a x x, 
      ∀x y z. P x y z ⟶ P (f x) y (f z)⟧
     ⊢ ∃z. P (f a) z (f (f a))
  ------------------------------------------------------------------ *}

lemma ejercicio_34a:
  assumes "∀x. P a x x" 
          "∀x y z. P x y z ⟶ P (f x) y (f z)"
  shows   "∃z. P (f a) z (f (f a))"
  using assms by metis

lemma ejercicio_34b:
  assumes "∀x. P a x x" 
           "∀x y z. P x y z ⟶ P (f x) y (f z)"
  shows   "∃z. P (f a) z (f (f a))"
proof -
  have "P a (f a) (f a)" using assms(1) ..
  have "∀y z. P a y z ⟶ P (f a) y (f z)" using assms(2) ..
  hence "∀z. P a (f a) z ⟶ P (f a) (f a) (f z)" ..
  hence "P a (f a) (f a) ⟶ P (f a) (f a) (f (f a))" ..
  hence "P (f a) (f a) (f (f a))" using `P a (f a) (f a)` ..
  thus "∃z. P (f a) z (f (f a))" ..
qed

text {* --------------------------------------------------------------- 
  Ejercicio 35. Demostrar o refutar
     {∀y. Q a y, 
      ∀x y. Q x y ⟶ Q (s x) (s y)} 
     ⊢ ∃z. Qa z ∧ Q z (s (s a))
  ------------------------------------------------------------------ *}

lemma ejercicio_35a:
  assumes "∀y. Q a y" 
          "∀x y. Q x y ⟶ Q (s x) (s y)" 
  shows   "∃z. Q a z ∧ Q z (s (s a))"
  using assms by metis

lemma ejercicio_35b:
  assumes "∀y. Q a y" 
          "∀x y. Q x y ⟶ Q (s x) (s y)" 
  shows   "∃z. Q a z ∧ Q z (s (s a))"
proof -
  have "Q a (s a)" using assms(1) ..
  have "∀y. Q a y ⟶ Q (s a) (s y)" using assms(2) ..
  hence "Q a (s a) ⟶ Q (s a) (s (s a))" ..
  hence "Q (s a) (s (s a))" using `Q a (s a)` ..
  with `Q a (s a)` have "Q a (s a) ∧ Q (s a) (s (s a))" ..
  thus "∃z. Q a z ∧ Q z (s (s a))" ..
qed

text {* --------------------------------------------------------------- 
  Ejercicio 36. Demostrar o refutar
     {x = f x, odd (f x)} ⊢ odd x
  ------------------------------------------------------------------ *}

lemma ejercicio_36a:
  "⟦x = f x; odd (f x)⟧ ⟹ odd x"
by auto


lemma ejercicio_36b:
  assumes "x = f x" and
          "odd (f x)"
  shows "odd x"
proof -
  show "odd x" using assms by (rule ssubst)
qed

text {* --------------------------------------------------------------- 
  Ejercicio 37. Demostrar o refutar
     {x = f x, triple (f x) (f x) x} ⊢ triple x x x
  ------------------------------------------------------------------ *}

lemma ejercicio_37a:
  "⟦x = f x; triple (f x) (f x) x⟧ ⟹ triple x x x"
  by auto

lemma ejercicio_37b:
  assumes "x = f x" and
          "triple (f x) (f x) x"
  shows "triple x x x"
proof -
  show "triple x x x" using assms by (rule ssubst)
qed

lemma ej_8_8_26: 
  assumes 1:"∀x. P x ⟶ ((∃y. Q x y) ⟶ (∃y. Q y x)) " and
          2:"∀x. (∃y. Q y x) ⟶ Q x x" and
          3: "¬ (∃x. Q x x)"
  shows   "∀x. P x ⟶ (∀y. (¬ Q x y))"
proof
  fix a 
  show "P a ⟶ (∀y. (¬ Q a y))"
   proof
      assume "P a"
        show "∀y. (¬ Q a y)"
         proof
            fix b
            show "¬ Q a b"
              proof
                 assume "Q a b"
                 hence 4:"∃y. Q a y" ..
                 have "P a ⟶ ((∃y. Q a y) ⟶ (∃y. Q y a))" using 1 ..
                 hence "(∃y. Q a y) ⟶ (∃y. Q y a)" using `P a` ..
                 hence 5:"∃y. Q y a" using 4 ..
                 have "(∃y. Q y a) ⟶ Q a a" using 2 ..
                 hence "Q a a" using 5 ..
                 hence "∃x. Q x x" ..
                 with 3 show False ..
              qed
         qed
   qed
qed

end