(* Title: HOL/ex/Refute_Examples.thy ID: $Id: Refute_Examples.thy,v 1.12 2005/07/26 10:13:35 webertj Exp $ Author: Tjark Weber Copyright 2003-2005 *) (* See 'HOL/Refute.thy' for help. *) header {* Examples for the 'refute' command *} theory Refute_Examples imports Main begin lemma "P ∧ Q" apply (rule conjI) refute 1 -- {* refutes @{term "P"} *} refute 2 -- {* refutes @{term "Q"} *} refute -- {* equivalent to 'refute 1' *} -- {* here 'refute 3' would cause an exception, since we only have 2 subgoals *} refute [maxsize=5] -- {* we can override parameters ... *} refute [satsolver="dpll"] 2 -- {* ... and specify a subgoal at the same time *} oops section {* Examples and Test Cases *} subsection {* Propositional logic *} lemma "True" refute apply auto done lemma "False" refute oops lemma "P" refute oops lemma "~ P" refute oops lemma "P & Q" refute oops lemma "P | Q" refute oops lemma "P --> Q" refute oops lemma "(P::bool) = Q" refute oops lemma "(P | Q) --> (P & Q)" refute oops subsection {* Predicate logic *} lemma "P x y z" refute oops lemma "P x y --> P y x" refute oops lemma "P (f (f x)) --> P x --> P (f x)" refute oops subsection {* Equality *} lemma "P = True" refute oops lemma "P = False" refute oops lemma "x = y" refute oops lemma "f x = g x" refute oops lemma "(f::'a=>'b) = g" refute oops lemma "(f::('d=>'d)=>('c=>'d)) = g" refute oops lemma "distinct [a,b]" refute apply simp refute oops subsection {* First-Order Logic *} lemma "∃x. P x" refute oops lemma "∀x. P x" refute oops lemma "EX! x. P x" refute oops lemma "Ex P" refute oops lemma "All P" refute oops lemma "Ex1 P" refute oops lemma "(∃x. P x) --> (∀x. P x)" refute oops lemma "(∀x. ∃y. P x y) --> (∃y. ∀x. P x y)" refute oops lemma "(∃x. P x) --> (EX! x. P x)" refute oops text {* A true statement (also testing names of free and bound variables being identical) *} lemma "(∀x y. P x y --> P y x) --> (∀x. P x y) --> P y x" refute apply fast done text {* "A type has at most 5 elements." *} lemma "a=b | a=c | a=d | a=e | a=f | b=c | b=d | b=e | b=f | c=d | c=e | c=f | d=e | d=f | e=f" refute oops lemma "∀a b c d e f. a=b | a=c | a=d | a=e | a=f | b=c | b=d | b=e | b=f | c=d | c=e | c=f | d=e | d=f | e=f" refute -- {* quantification causes an expansion of the formula; the previous version with free variables is refuted much faster *} oops text {* "Every reflexive and symmetric relation is transitive." *} lemma "[| ∀x. P x x; ∀x y. P x y --> P y x |] ==> P x y --> P y z --> P x z" refute oops text {* The "Drinker's theorem" ... *} lemma "∃x. f x = g x --> f = g" refute [maxsize=4] apply (auto simp add: ext) done text {* ... and an incorrect version of it *} lemma "(∃x. f x = g x) --> f = g" refute oops text {* "Every function has a fixed point." *} lemma "∃x. f x = x" refute oops text {* "Function composition is commutative." *} lemma "f (g x) = g (f x)" refute oops text {* "Two functions that are equivalent wrt.\ the same predicate 'P' are equal." *} lemma "((P::('a=>'b)=>bool) f = P g) --> (f x = g x)" refute oops subsection {* Higher-Order Logic *} lemma "∃P. P" refute apply auto done lemma "∀P. P" refute oops lemma "EX! P. P" refute apply auto done lemma "EX! P. P x" refute oops lemma "P Q | Q x" refute oops lemma "P All" refute oops lemma "P Ex" refute oops lemma "P Ex1" refute oops text {* "The transitive closure 'T' of an arbitrary relation 'P' is non-empty." *} constdefs "trans" :: "('a => 'a => bool) => bool" "trans P == (ALL x y z. P x y --> P y z --> P x z)" "subset" :: "('a => 'a => bool) => ('a => 'a => bool) => bool" "subset P Q == (ALL x y. P x y --> Q x y)" "trans_closure" :: "('a => 'a => bool) => ('a => 'a => bool) => bool" "trans_closure P Q == (subset Q P) & (trans P) & (ALL R. subset Q R --> trans R --> subset P R)" lemma "trans_closure T P --> (∃x y. T x y)" refute oops text {* "The union of transitive closures is equal to the transitive closure of unions." *} lemma "(∀x y. (P x y | R x y) --> T x y) --> trans T --> (∀Q. (∀x y. (P x y | R x y) --> Q x y) --> trans Q --> subset T Q) --> trans_closure TP P --> trans_closure TR R --> (T x y = (TP x y | TR x y))" refute oops text {* "Every surjective function is invertible." *} lemma "(∀y. ∃x. y = f x) --> (∃g. ∀x. g (f x) = x)" refute oops text {* "Every invertible function is surjective." *} lemma "(∃g. ∀x. g (f x) = x) --> (∀y. ∃x. y = f x)" refute oops text {* Every point is a fixed point of some function. *} lemma "∃f. f x = x" refute [maxsize=4] apply (rule_tac x="λx. x" in exI) apply simp done text {* Axiom of Choice: first an incorrect version ... *} lemma "(∀x. ∃y. P x y) --> (EX!f. ∀x. P x (f x))" refute oops text {* ... and now two correct ones *} lemma "(∀x. ∃y. P x y) --> (∃f. ∀x. P x (f x))" refute [maxsize=4] apply (simp add: choice) done lemma "(∀x. EX!y. P x y) --> (EX!f. ∀x. P x (f x))" refute [maxsize=2] apply auto apply (simp add: ex1_implies_ex choice) apply (fast intro: ext) done subsection {* Meta-logic *} lemma "!!x. P x" refute oops lemma "f x == g x" refute oops lemma "P ==> Q" refute oops lemma "[| P; Q; R |] ==> S" refute oops subsection {* Schematic variables *} lemma "?P" refute apply auto done lemma "x = ?y" refute apply auto done subsection {* Abstractions *} lemma "(λx. x) = (λx. y)" refute oops lemma "(λf. f x) = (λf. True)" refute oops lemma "(λx. x) = (λy. y)" refute apply simp done subsection {* Sets *} lemma "P (A::'a set)" refute oops lemma "P (A::'a set set)" refute oops lemma "{x. P x} = {y. P y}" refute apply simp done lemma "x : {x. P x}" refute oops lemma "P op:" refute oops lemma "P (op: x)" refute oops lemma "P Collect" refute oops lemma "A Un B = A Int B" refute oops lemma "(A Int B) Un C = (A Un C) Int B" refute oops lemma "Ball A P --> Bex A P" refute oops subsection {* arbitrary *} lemma "arbitrary" refute oops lemma "P arbitrary" refute oops lemma "arbitrary x" refute oops lemma "arbitrary arbitrary" refute oops subsection {* The *} lemma "The P" refute oops lemma "P The" refute oops lemma "P (The P)" refute oops lemma "(THE x. x=y) = z" refute oops lemma "Ex P --> P (The P)" refute oops subsection {* Eps *} lemma "Eps P" refute oops lemma "P Eps" refute oops lemma "P (Eps P)" refute oops lemma "(SOME x. x=y) = z" refute oops lemma "Ex P --> P (Eps P)" refute [maxsize=3] apply (auto simp add: someI) done (******************************************************************************) subsection {* Subtypes (typedef), typedecl *} text {* A completely unspecified non-empty subset of @{typ "'a"}: *} typedef 'a myTdef = "insert (arbitrary::'a) (arbitrary::'a set)" by auto lemma "(x::'a myTdef) = y" refute oops typedecl myTdecl typedef 'a T_bij = "{(f::'a=>'a). ∀y. ∃!x. f x = y}" by auto lemma "P (f::(myTdecl myTdef) T_bij)" refute oops (******************************************************************************) subsection {* Inductive datatypes *} text {* With quick\_and\_dirty set, the datatype package does not generate certain axioms for recursion operators. Without these axioms, refute may find spurious countermodels. *} ML {* reset quick_and_dirty; *} subsubsection {* unit *} lemma "P (x::unit)" refute oops lemma "∀x::unit. P x" refute oops lemma "P ()" refute oops lemma "P (unit_rec u x)" refute oops lemma "P (case x of () => u)" refute oops subsubsection {* option *} lemma "P (x::'a option)" refute oops lemma "∀x::'a option. P x" refute oops lemma "P None" refute oops lemma "P (Some x)" refute oops lemma "P (option_rec n s x)" refute oops lemma "P (case x of None => n | Some u => s u)" refute oops subsubsection {* * *} lemma "P (x::'a*'b)" refute oops lemma "∀x::'a*'b. P x" refute oops lemma "P (x,y)" refute oops lemma "P (fst x)" refute oops lemma "P (snd x)" refute oops lemma "P Pair" refute oops lemma "P (prod_rec p x)" refute oops lemma "P (case x of Pair a b => p a b)" refute oops subsubsection {* + *} lemma "P (x::'a+'b)" refute oops lemma "∀x::'a+'b. P x" refute oops lemma "P (Inl x)" refute oops lemma "P (Inr x)" refute oops lemma "P Inl" refute oops lemma "P (sum_rec l r x)" refute oops lemma "P (case x of Inl a => l a | Inr b => r b)" refute oops subsubsection {* Non-recursive datatypes *} datatype T1 = A | B lemma "P (x::T1)" refute oops lemma "∀x::T1. P x" refute oops lemma "P A" refute oops lemma "P (T1_rec a b x)" refute oops lemma "P (case x of A => a | B => b)" refute oops datatype 'a T2 = C T1 | D 'a lemma "P (x::'a T2)" refute oops lemma "∀x::'a T2. P x" refute oops lemma "P D" refute oops lemma "P (T2_rec c d x)" refute oops lemma "P (case x of C u => c u | D v => d v)" refute oops datatype ('a,'b) T3 = E "'a => 'b" lemma "P (x::('a,'b) T3)" refute oops lemma "∀x::('a,'b) T3. P x" refute oops lemma "P E" refute oops lemma "P (T3_rec e x)" refute oops lemma "P (case x of E f => e f)" refute oops subsubsection {* Recursive datatypes *} text {* nat *} lemma "P (x::nat)" refute oops lemma "∀x::nat. P x" refute oops lemma "P (Suc 0)" refute oops lemma "P Suc" refute -- {* @{term "Suc"} is a partial function (regardless of the size of the model), hence @{term "P Suc"} is undefined, hence no model will be found *} oops lemma "P (nat_rec zero suc x)" refute oops lemma "P (case x of 0 => zero | Suc n => suc n)" refute oops text {* 'a list *} lemma "P (xs::'a list)" refute oops lemma "∀xs::'a list. P xs" refute oops lemma "P [x, y]" refute oops lemma "P (list_rec nil cons xs)" refute oops lemma "P (case x of Nil => nil | Cons a b => cons a b)" refute oops lemma "(xs::'a list) = ys" refute oops lemma "a # xs = b # xs" refute oops datatype 'a BinTree = Leaf 'a | Node "'a BinTree" "'a BinTree" lemma "P (x::'a BinTree)" refute oops lemma "∀x::'a BinTree. P x" refute oops lemma "P (Node (Leaf x) (Leaf y))" refute oops lemma "P (BinTree_rec l n x)" refute oops lemma "P (case x of Leaf a => l a | Node a b => n a b)" refute oops subsubsection {* Mutually recursive datatypes *} datatype 'a aexp = Number 'a | ITE "'a bexp" "'a aexp" "'a aexp" and 'a bexp = Equal "'a aexp" "'a aexp" lemma "P (x::'a aexp)" refute oops lemma "∀x::'a aexp. P x" refute oops lemma "P (ITE (Equal (Number x) (Number y)) (Number x) (Number y))" refute oops lemma "P (x::'a bexp)" refute oops lemma "∀x::'a bexp. P x" refute oops lemma "P (aexp_bexp_rec_1 number ite equal x)" refute oops lemma "P (case x of Number a => number a | ITE b a1 a2 => ite b a1 a2)" refute oops lemma "P (aexp_bexp_rec_2 number ite equal x)" refute oops lemma "P (case x of Equal a1 a2 => equal a1 a2)" refute oops subsubsection {* Other datatype examples *} datatype Trie = TR "Trie list" lemma "P (x::Trie)" refute oops lemma "∀x::Trie. P x" refute oops lemma "P (TR [TR []])" refute oops lemma "P (Trie_rec_1 a b c x)" refute oops lemma "P (Trie_rec_2 a b c x)" refute oops datatype InfTree = Leaf | Node "nat => InfTree" lemma "P (x::InfTree)" refute oops lemma "∀x::InfTree. P x" refute oops lemma "P (Node (λn. Leaf))" refute oops lemma "P (InfTree_rec leaf node x)" refute oops datatype 'a lambda = Var 'a | App "'a lambda" "'a lambda" | Lam "'a => 'a lambda" lemma "P (x::'a lambda)" refute oops lemma "∀x::'a lambda. P x" refute oops lemma "P (Lam (λa. Var a))" refute oops lemma "P (lambda_rec v a l x)" refute oops text {* Taken from "Inductive datatypes in HOL", p.8: *} datatype ('a, 'b) T = C "'a => bool" | D "'b list" datatype 'c U = E "('c, 'c U) T" lemma "P (x::'c U)" refute oops lemma "∀x::'c U. P x" refute oops lemma "P (E (C (λa. True)))" refute oops lemma "P (U_rec_1 e f g h i x)" refute oops lemma "P (U_rec_2 e f g h i x)" refute oops lemma "P (U_rec_3 e f g h i x)" refute oops (******************************************************************************) subsection {* Records *} (*TODO: make use of pair types, rather than typedef, for record types*) record ('a, 'b) point = xpos :: 'a ypos :: 'b lemma "(x::('a, 'b) point) = y" refute oops record ('a, 'b, 'c) extpoint = "('a, 'b) point" + ext :: 'c lemma "(x::('a, 'b, 'c) extpoint) = y" refute oops (******************************************************************************) subsection {* Inductively defined sets *} consts arbitrarySet :: "'a set" inductive arbitrarySet intros "arbitrary : arbitrarySet" lemma "x : arbitrarySet" refute oops consts evenCard :: "'a set set" inductive evenCard intros "{} : evenCard" "[| S : evenCard; x ∉ S; y ∉ S; x ≠ y |] ==> S ∪ {x, y} : evenCard" lemma "S : evenCard" refute oops consts even :: "nat set" odd :: "nat set" inductive even odd intros "0 : even" "n : even ==> Suc n : odd" "n : odd ==> Suc n : even" lemma "n : odd" (*refute*) -- {* unfortunately, this little example already takes too long *} oops (******************************************************************************) subsection {* Examples involving special functions *} lemma "card x = 0" refute oops lemma "finite x" refute -- {* no finite countermodel exists *} oops lemma "(x::nat) + y = 0" refute oops lemma "(x::nat) = x + x" refute oops lemma "(x::nat) - y + y = x" refute oops lemma "(x::nat) = x * x" refute oops lemma "(x::nat) < x + y" refute oops lemma "a @ [] = b @ []" refute oops lemma "a @ b = b @ a" refute oops lemma "f (lfp f) = lfp f" refute oops lemma "f (gfp f) = gfp f" refute oops lemma "lfp f = gfp f" refute oops (******************************************************************************) subsection {* Axiomatic type classes and overloading *} text {* A type class without axioms: *} axclass classA lemma "P (x::'a::classA)" refute oops text {* The axiom of this type class does not contain any type variables, but is internally converted into one that does: *} axclass classB classB_ax: "P | ~ P" lemma "P (x::'a::classB)" refute oops text {* An axiom with a type variable (denoting types which have at least two elements): *} axclass classC < type classC_ax: "∃x y. x ≠ y" lemma "P (x::'a::classC)" refute oops lemma "∃x y. (x::'a::classC) ≠ y" refute -- {* no countermodel exists *} oops text {* A type class for which a constant is defined: *} consts classD_const :: "'a => 'a" axclass classD < type classD_ax: "classD_const (classD_const x) = classD_const x" lemma "P (x::'a::classD)" refute oops text {* A type class with multiple superclasses: *} axclass classE < classC, classD lemma "P (x::'a::classE)" refute oops lemma "P (x::'a::{classB, classE})" refute oops text {* OFCLASS: *} lemma "OFCLASS('a::type, type_class)" refute -- {* no countermodel exists *} apply intro_classes done lemma "OFCLASS('a::classC, type_class)" refute -- {* no countermodel exists *} apply intro_classes done lemma "OFCLASS('a, classB_class)" refute -- {* no countermodel exists *} apply intro_classes apply simp done lemma "OFCLASS('a::type, classC_class)" refute oops text {* Overloading: *} consts inverse :: "'a => 'a" defs (overloaded) inverse_bool: "inverse (b::bool) == ~ b" inverse_set : "inverse (S::'a set) == -S" inverse_pair: "inverse p == (inverse (fst p), inverse (snd p))" lemma "inverse b" refute oops lemma "P (inverse (S::'a set))" refute oops lemma "P (inverse (p::'a×'b))" refute oops end
lemma
True
lemma
(∀x y. P x y --> P y x) --> (∀x. P x y) --> P y x
lemma
∃x. f x = g x --> f = g
lemma
∃P. P
lemma
∃!P. P
lemma
∃f. f x = x
lemma
(∀x. ∃y. P x y) --> (∃f. ∀x. P x (f x))
lemma
(∀x. ∃!y. P x y) --> (∃!f. ∀x. P x (f x))
lemma
True
lemma
x = x
lemma
(%x. x) = (%y. y)
lemma
{x. P x} = {y. P y}
lemma
Ex P --> P (Eps P)
lemma
OFCLASS('a, type_class)
lemma
OFCLASS('a, type_class)
lemma
OFCLASS('a, classB_class)