thys/ProofAutomation.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Mon, 09 Feb 2015 12:13:10 +0000
changeset 62 a6bb0152ccc2
parent 29 2345ba5b4264
permissions -rw-r--r--
updated some rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     1
theory ProofAutomation
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     2
imports Main
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     3
begin
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     4
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     5
lemma "\<forall>x. \<exists>y. x = y"
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     6
by auto
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     7
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     8
lemma "A \<subseteq> B \<inter> C \<Longrightarrow> A \<subseteq> B \<union> C"
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
     9
by auto
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    10
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    11
lemma "\<lbrakk> \<forall> xs \<in> A. \<exists> ys. xs = ys @ ys; us \<in> A \<rbrakk> \<Longrightarrow> \<exists>n. length us = n+n"
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    12
by fastforce
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    13
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    14
lemma "\<lbrakk>xs @ ys = ys @ xs; length xs = length ys \<rbrakk> \<Longrightarrow> xs = ys"
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    15
by auto
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    16
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    17
lemma "\<lbrakk> (a::nat) \<le> x + b; 2*x < c\<rbrakk> \<Longrightarrow> 2*a + 1 \<le> 2*b + c"
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    18
by arith
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    19
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    20
lemma "\<lbrakk> (a::nat) \<le> b; b \<le> c; c \<le> d; d \<le> e \<rbrakk> \<Longrightarrow> a \<le> e"
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    21
by (blast intro: le_trans)
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    22
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    23
lemma "Suc(Suc(Suc a)) \<le> b \<Longrightarrow> a \<le> b"
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    24
by(blast dest: Suc_leD)
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    25
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    26
inductive ev :: "nat \<Rightarrow> bool" where
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    27
ev0: "ev 0" |
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    28
evSS: "ev n \<Longrightarrow> ev (n + 2)"
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    29
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    30
fun even :: "nat \<Rightarrow> bool" where
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    31
"even 0 = True" |
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    32
"even (Suc 0) = False" |
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    33
"even (Suc(Suc n)) = even n"
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    34
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    35
lemma "ev m \<Longrightarrow> even m"
2345ba5b4264 Proof Automation
fahadausaf <fahad.ausaf@icloud.com>
parents:
diff changeset
    36