Theory Fol1

theory Fol1
imports Trancl
(* 
    This file is a part of IsarMathLib - 
    a library of formalized mathematics for Isabelle/Isar.

    Copyright (C) 2005 - 2008  Slawomir Kolodynski

    This program is free software; Redistribution and use in source and binary forms, 
    with or without modification, are permitted provided that the following conditions are met:

   1. Redistributions of source code must retain the above copyright notice, 
   this list of conditions and the following disclaimer.
   2. Redistributions in binary form must reproduce the above copyright notice, 
   this list of conditions and the following disclaimer in the documentation and/or 
   other materials provided with the distribution.
   3. The name of the author may not be used to endorse or promote products 
   derived from this software without specific prior written permission.

THIS SOFTWARE IS PROVIDED BY THE AUTHOR ``AS IS'' AND ANY EXPRESS OR IMPLIED 
WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED WARRANTIES OF 
MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE DISCLAIMED. 
IN NO EVENT SHALL THE AUTHOR BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, 
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO, 
PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; 
OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, 
WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR 
OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, 
EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.

*)

section ‹First Order Logic›

theory Fol1 imports ZF.Trancl

begin

text‹Isabelle/ZF builds on the first order logic. Almost everything
  one would like to have in this area is covered in the standard Isabelle 
  libraries. The material in this theory provides some lemmas that are
  missing or allow for a more readable proof style.›


subsection‹Notions and lemmas in FOL›

text‹This section contains mostly shortcuts and workarounds 
  that allow to use more readable coding style.›

text‹The next lemma serves as a workaround to problems with applying 
  the definition of transitivity (of a relation) in our coding style 
  (any attempt to do
  something like ‹using trans_def› puts Isabelle in an 
  infinite loop).›

lemma Fol1_L2: assumes 
  A1: "∀ x y z. ⟨x, y⟩ ∈ r ∧ ⟨y, z⟩ ∈ r ⟶ ⟨x, z⟩ ∈ r"
  shows "trans(r)"
proof -
  from A1 have
    "∀ x y z. ⟨x, y⟩ ∈ r ⟶ ⟨y, z⟩ ∈ r ⟶ ⟨x, z⟩ ∈ r"
    using imp_conj by blast
  then show ?thesis unfolding trans_def by blast
qed

text‹Another workaround for the problem of Isabelle simplifier looping when 
  the transitivity definition is used.›

lemma Fol1_L3: assumes A1: "trans(r)" and A2: "⟨ a,b⟩ ∈ r  ∧ ⟨ b,c⟩ ∈ r"
  shows "⟨ a,c⟩ ∈ r"
proof -
  from A1 have  "∀x y z. ⟨x, y⟩ ∈ r ⟶ ⟨y, z⟩ ∈ r ⟶ ⟨x, z⟩ ∈ r"
   unfolding trans_def by blast
  with A2 show ?thesis using imp_conj by fast
qed
  
text‹There is a problem with application of the definition of asymetry for
  relations. The next lemma is a workaround.›

lemma Fol1_L4: 
  assumes A1: "antisym(r)" and A2: "⟨ a,b⟩ ∈ r"   "⟨ b,a⟩ ∈ r"  
  shows "a=b"
proof -
  from A1 have "∀ x y. ⟨ x,y⟩ ∈ r ⟶ ⟨ y,x⟩ ∈ r ⟶ x=y"
    unfolding antisym_def by blast
  with A2 show "a=b" using imp_conj by fast
qed

text‹The definition below implements a common idiom that states that 
  (perhaps under some assumptions) exactly one of given three statements 
  is true.›

definition
  "Exactly_1_of_3_holds(p,q,r) ≡ 
  (p∨q∨r) ∧ (p ⟶ ¬q ∧ ¬r) ∧ (q ⟶ ¬p ∧ ¬r) ∧ (r ⟶ ¬p ∧ ¬q)"

text‹The next lemma allows to prove statements of the form 
  ‹Exactly_1_of_3_holds(p,q,r)›.›

lemma Fol1_L5:
  assumes "p∨q∨r"
  and "p ⟶ ¬q ∧ ¬r"
  and "q ⟶ ¬p ∧ ¬r"
  and "r ⟶ ¬p ∧ ¬q"
  shows "Exactly_1_of_3_holds(p,q,r)"
proof -
  from assms have
    "(p∨q∨r) ∧ (p ⟶ ¬q ∧ ¬r) ∧ (q ⟶ ¬p ∧ ¬r) ∧ (r ⟶ ¬p ∧ ¬q)"
    by blast
  then show "Exactly_1_of_3_holds (p,q,r)"
    unfolding Exactly_1_of_3_holds_def by fast
qed

text‹If exactly one of $p,q,r$ holds and $p$ is not true, then
  $q$ or $r$.›

lemma Fol1_L6: 
  assumes A1: "¬p" and A2: "Exactly_1_of_3_holds(p,q,r)" 
  shows "q∨r"
proof -
  from A2 have  
    "(p∨q∨r) ∧ (p ⟶ ¬q ∧ ¬r) ∧ (q ⟶ ¬p ∧ ¬r) ∧ (r ⟶ ¬p ∧ ¬q)"
    unfolding Exactly_1_of_3_holds_def by fast
  hence "p ∨ q ∨ r" by blast
  with A1 show "q ∨ r" by simp
qed

text‹If exactly one of $p,q,r$ holds and $q$ is true, then 
  $r$ can not be true.›

lemma Fol1_L7:
  assumes A1: "q" and A2: "Exactly_1_of_3_holds(p,q,r)"
  shows "¬r"
proof -
   from A2 have  
    "(p∨q∨r) ∧ (p ⟶ ¬q ∧ ¬r) ∧ (q ⟶ ¬p ∧ ¬r) ∧ (r ⟶ ¬p ∧ ¬q)"
    unfolding Exactly_1_of_3_holds_def by fast
  with A1 show "¬r" by blast
qed

text‹The next lemma demonstrates an elegant form of the 
  ‹Exactly_1_of_3_holds(p,q,r)› predicate.›

lemma Fol1_L8: 
  shows "Exactly_1_of_3_holds(p,q,r) ⟷ (p⟷q⟷r) ∧ ¬(p∧q∧r)"
proof
  assume "Exactly_1_of_3_holds(p,q,r)"
  then have 
    "(p∨q∨r) ∧ (p ⟶ ¬q ∧ ¬r) ∧ (q ⟶ ¬p ∧ ¬r) ∧ (r ⟶ ¬p ∧ ¬q)"
    unfolding Exactly_1_of_3_holds_def by fast
  thus "(p⟷q⟷r) ∧ ¬(p∧q∧r)" by blast
next assume "(p⟷q⟷r) ∧ ¬(p∧q∧r)" 
  hence
    "(p∨q∨r) ∧ (p ⟶ ¬q ∧ ¬r) ∧ (q ⟶ ¬p ∧ ¬r) ∧ (r ⟶ ¬p ∧ ¬q)"
    by auto
  then show "Exactly_1_of_3_holds(p,q,r)"
    unfolding Exactly_1_of_3_holds_def by fast
qed

text‹A property of the ‹Exactly_1_of_3_holds› predicate.›

lemma Fol1_L8A: assumes A1: "Exactly_1_of_3_holds(p,q,r)"
  shows "p ⟷ ¬(q ∨ r)"
proof -
  from A1 have "(p∨q∨r) ∧ (p ⟶ ¬q ∧ ¬r) ∧ (q ⟶ ¬p ∧ ¬r) ∧ (r ⟶ ¬p ∧ ¬q)"
    unfolding Exactly_1_of_3_holds_def by fast
  then show "p ⟷ ¬(q ∨ r)" by blast
qed

text‹Exclusive or definition. There is one also defined in the standard 
  Isabelle, denoted ‹xor›, but it relates to boolean values, 
  which are sets. Here we define a logical functor.›

definition
  Xor (infixl "Xor" 66) where
  "p Xor q ≡ (p∨q) ∧ ¬(p ∧ q)"

text‹The "exclusive or" is the same as negation of equivalence.›

lemma Fol1_L9: shows "p Xor q ⟷ ¬(p⟷q)"
  using Xor_def by auto

text‹Equivalence relations are symmetric.›

lemma equiv_is_sym: assumes A1: "equiv(X,r)" and A2: "⟨x,y⟩ ∈ r"
  shows  "⟨y,x⟩ ∈ r"
proof -
  from A1 have "sym(r)" using equiv_def by simp
  then have "∀x y. ⟨x,y⟩ ∈ r ⟶ ⟨y,x⟩ ∈ r"
    unfolding sym_def by fast
  with A2 show "⟨y,x⟩ ∈ r" by blast
qed

(* In Isabelle/ZF conjunction associates to the right!.
lemma test: assumes A1: "P" "Q∧R"
  shows "P∧Q∧R"
proof - 
  from A1 show "P∧Q∧R" by (rule one_more_conj);
qed;

*)
end