Acciones

R9

De Razonamiento automático (2019-20)

chapter 9: Deducción natural de primer orden

theory R9_Deduccion_natural_de_primer_orden
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

  · allI:       ⟦∀x. P x; P x  R  R
  · allE:       (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_1: 
  assumes "∀x. P x ⟶ Q x"
  shows   "(∀x. P x) ⟶ (∀x. Q x)"
oops

text --------------------------------------------------------------- 
  Ejercicio 2. Demostrar
       x. ¬(P x)  ¬(x. P x)
  ------------------------------------------------------------------

lemma ejercicio_2: 
  assumes "∃x. ¬(P x)"
  shows   "¬(∀x. P x)"
oops

text --------------------------------------------------------------- 
  Ejercicio 3. Demostrar
       x. P x  y. P y
  ------------------------------------------------------------------

lemma ejercicio_3: 
  assumes "∀x. P x"
  shows   "∀y. P y"
oops

text --------------------------------------------------------------- 
  Ejercicio 4. Demostrar
       x. P x  Q x  (x. ¬(Q x))  (x. ¬ (P x))
  ------------------------------------------------------------------

lemma ejercicio_4: 
  assumes "∀x. P x ⟶ Q x"
  shows   "(∀x. ¬(Q x)) ⟶ (∀x. ¬ (P x))"
oops

text --------------------------------------------------------------- 
  Ejercicio 5. Demostrar
       x. P x   ¬(Q x)  ¬(x. P x  Q x)
  ------------------------------------------------------------------

lemma ejercicio_5: 
  assumes "∀x. P x  ⟶ ¬(Q x)"
  shows   "¬(∃x. P x ∧ Q x)"
oops

text --------------------------------------------------------------- 
  Ejercicio 6. Demostrar
       x y. P x y  u v. P u v
  ------------------------------------------------------------------

lemma ejercicio_6: 
  assumes "∀x y. P x y"
  shows   "∀u v. P u v"
oops

text --------------------------------------------------------------- 
  Ejercicio 7. Demostrar
       x y. P x y  u v. P u v
  ------------------------------------------------------------------

lemma ejercicio_7: 
  assumes "∃x y. P x y"
  shows   "∃u v. P u v"
oops

text --------------------------------------------------------------- 
  Ejercicio 8. Demostrar
       x. y. P x y  y. x. P x y
  ------------------------------------------------------------------

lemma ejercicio_8: 
  assumes "∃x. ∀y. P x y"
  shows   "∀y. ∃x. P x y"
oops

text --------------------------------------------------------------- 
  Ejercicio 9. Demostrar
       x. P a  Q x  P a  (x. Q x)
  ------------------------------------------------------------------

lemma ejercicio_9: 
  assumes "∃x. P a ⟶ Q x"
  shows   "P a ⟶ (∃x. Q x)"
oops

text --------------------------------------------------------------- 
  Ejercicio 10. Demostrar
       P a  (x. Q x)  x. P a  Q x 
  ------------------------------------------------------------------

lemma ejercicio_10: 
  fixes P Q :: "'b ⇒ bool" 
  assumes "P a ⟶ (∃x. Q x)"
  shows   "∃x. P a ⟶ Q x"
oops

text --------------------------------------------------------------- 
  Ejercicio 11. Demostrar
       (x. P x)  Q a  x. P x  Q a
  ------------------------------------------------------------------

lemma ejercicio_11: 
  assumes "(∃x. P x) ⟶ Q a"
  shows   "∀x. P x ⟶ Q a"
oops

text --------------------------------------------------------------- 
  Ejercicio 12. Demostrar
       x. P x  Q a   x. P x  Q a
  ------------------------------------------------------------------

lemma ejercicio_12: 
  assumes "∀x. P x ⟶ Q a"
  shows   "∃x. P x ⟶ Q a"
oops

text --------------------------------------------------------------- 
  Ejercicio 13. Demostrar
       (x. P x)  (x. Q x)  x. P x  Q x
  ------------------------------------------------------------------

lemma ejercicio_13: 
  assumes "(∀x. P x) ∨ (∀x. Q x)"
  shows   "∀x. P x ∨ Q x"
oops

text --------------------------------------------------------------- 
  Ejercicio 14. Demostrar
       x. P x  Q x  (x. P x)  (x. Q x)
  ------------------------------------------------------------------

lemma ejercicio_14: 
  assumes "∃x. P x ∧ Q x"
  shows   "(∃x. P x) ∧ (∃x. Q x)"
oops

text --------------------------------------------------------------- 
  Ejercicio 15. Demostrar
       x y. P y  Q x  (y. P y)  (x. Q x)
  ------------------------------------------------------------------

lemma ejercicio_15: 
  assumes "∀x y. P y ⟶ Q x"
  shows   "(∃y. P y) ⟶ (∀x. Q x)"
oops

text --------------------------------------------------------------- 
  Ejercicio 16. Demostrar
       ¬(x. ¬(P x))  x. P x
  ------------------------------------------------------------------

lemma ejercicio_16: 
  assumes "¬(∀x. ¬(P x))"
  shows   "∃x. P x"
oops

text --------------------------------------------------------------- 
  Ejercicio 17. Demostrar
       x. ¬(P x)  ¬(x. P x)
  ------------------------------------------------------------------

lemma ejercicio_17: 
  assumes "∀x. ¬(P x)"
  shows   "¬(∃x. P x)"
oops

text --------------------------------------------------------------- 
  Ejercicio 18. Demostrar
       x. P x  ¬(x. ¬(P x))
  ------------------------------------------------------------------

lemma ejercicio_18: 
  assumes "∃x. P x"
  shows   "¬(∀x. ¬(P x))"
oops

text --------------------------------------------------------------- 
  Ejercicio 19. Demostrar
       P a  (x. Q x)  x. P a  Q x
  ------------------------------------------------------------------

lemma ejercicio_19: 
  assumes "P a ⟶ (∀x. Q x)"
  shows   "∀x. P a ⟶ Q x"
oops

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_20: 
  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)"
oops

text --------------------------------------------------------------- 
  Ejercicio 21. Demostrar
     {x. P x  Q x, x. ¬(Q x), x. R x  ¬(P x)}  x. ¬(R x)
  ------------------------------------------------------------------

lemma ejercicio_21:
  assumes "∀x. P x ∨ Q x" 
          "∃x. ¬(Q x)" 
          "∀x. R x ⟶ ¬(P x)"
  shows   "∃x. ¬(R x)" 
oops

text --------------------------------------------------------------- 
  Ejercicio 22. Demostrar
     {x. P x  Q x  R x, ¬(x. P x  R x)}  x. P x  Q x
  ------------------------------------------------------------------

lemma ejercicio_22:
  assumes "∀x. P x ⟶ Q x ∨ R x" 
          "¬(∃x. P x ∧ R x)"
  shows   "∀x. P x ⟶ Q x"
oops

text --------------------------------------------------------------- 
  Ejercicio 23. Demostrar
     x y. R x y  R y x  x y. R x y
  ------------------------------------------------------------------

lemma ejercicio_23:
  assumes "∃x y. R x y ∨ R y x"
  shows   "∃x y. R x y"
oops

text --------------------------------------------------------------- 
  Ejercicio 24. Demostrar
       (x. y. P x y)  (y. x. P x y)
  ------------------------------------------------------------------

lemma ejercicio_24: 
  "(∃x. ∀y. P x y) ⟶ (∀y. ∃x. P x y)"
oops

text --------------------------------------------------------------- 
  Ejercicio 25. Demostrar
       (x. P x  Q)  ((x. P x)  Q)
  ------------------------------------------------------------------

lemma ejercicio_25: 
  "(∀x. P x ⟶ Q) ⟷ ((∃x. P x) ⟶ Q)"
oops

text --------------------------------------------------------------- 
  Ejercicio 26. Demostrar
       ((x. P x)  (x. Q x))  (x. P x  Q x)
  ------------------------------------------------------------------

lemma ejercicio_26: 
  "((∀x. P x) ∧ (∀x. Q x)) ⟷ (∀x. P x ∧ Q x)"
oops

text --------------------------------------------------------------- 
  Ejercicio 27. Demostrar o refutar
       ((x. P x)  (x. Q x))  (x. P x  Q x)
  ------------------------------------------------------------------

lemma ejercicio_27: 
  "((∀x. P x) ∨ (∀x. Q x)) ⟷ (∀x. P x ∨ Q x)"
oops

text --------------------------------------------------------------- 
  Ejercicio 28. Demostrar o refutar
       ((x. P x)  (x. Q x))  (x. P x  Q x)
  ------------------------------------------------------------------

lemma ejercicio_28: 
  "((∃x. P x) ∨ (∃x. Q x)) ⟷ (∃x. P x ∨ Q x)"
oops

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)"
oops

text --------------------------------------------------------------- 
  Ejercicio 30. Demostrar o refutar
       (¬(x. P x))  (x. ¬P x)
  ------------------------------------------------------------------

lemma ejercicio_30: 
  "(¬(∀x. P x)) ⟷ (∃x. ¬P x)"
oops

section Ejercicios sobre igualdad

text --------------------------------------------------------------- 
  Ejercicio 31. Demostrar o refutar
       P a  x. x = a  P x
  ------------------------------------------------------------------

lemma ejercicio_31b:
  assumes "P a"
  shows   "∀x. x = a ⟶ P x"
oops

text --------------------------------------------------------------- 
  Ejercicio 32. Demostrar o refutar
       x y. R x y  R y x; ¬(x. R x x)  x y. x  y
  ------------------------------------------------------------------

lemma ejercicio_32:
  fixes R :: "'c ⇒ 'c ⇒ bool"
  assumes "∃x y. R x y ∨ R y x"
          "¬(∃x. R x x)"
  shows   "∃(x::'c) y. x ≠ y"
oops

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_33:
  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)"
oops

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_34:
  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))"
oops

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_35:
  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))"
oops

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

lemma ejercicio_36:
  "⟦x = f x; odd (f x)⟧ ⟹ odd x"
oops

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

lemma ejercicio_37:
  "⟦x = f x; triple (f x) (f x) x⟧ ⟹ triple x x x"
oops

end