72029
|
1 |
(* Title: HOL/Examples/Coherent.thy
|
|
2 |
Author: Stefan Berghofer, TU Muenchen
|
|
3 |
Author: Marc Bezem, Institutt for Informatikk, Universitetet i Bergen
|
|
4 |
*)
|
|
5 |
|
|
6 |
section \<open>Coherent Logic Problems\<close>
|
|
7 |
|
|
8 |
theory Coherent
|
|
9 |
imports Main
|
|
10 |
begin
|
|
11 |
|
|
12 |
subsection \<open>Equivalence of two versions of Pappus' Axiom\<close>
|
|
13 |
|
81134
|
14 |
no_notation comp (infixl \<open>o\<close> 55)
|
|
15 |
unbundle no relcomp_syntax
|
72029
|
16 |
|
|
17 |
lemma p1p2:
|
|
18 |
assumes "col a b c l \<and> col d e f m"
|
|
19 |
and "col b f g n \<and> col c e g o"
|
|
20 |
and "col b d h p \<and> col a e h q"
|
|
21 |
and "col c d i r \<and> col a f i s"
|
|
22 |
and "el n o \<Longrightarrow> goal"
|
|
23 |
and "el p q \<Longrightarrow> goal"
|
|
24 |
and "el s r \<Longrightarrow> goal"
|
|
25 |
and "\<And>A. el A A \<Longrightarrow> pl g A \<Longrightarrow> pl h A \<Longrightarrow> pl i A \<Longrightarrow> goal"
|
|
26 |
and "\<And>A B C D. col A B C D \<Longrightarrow> pl A D"
|
|
27 |
and "\<And>A B C D. col A B C D \<Longrightarrow> pl B D"
|
|
28 |
and "\<And>A B C D. col A B C D \<Longrightarrow> pl C D"
|
|
29 |
and "\<And>A B. pl A B \<Longrightarrow> ep A A"
|
|
30 |
and "\<And>A B. ep A B \<Longrightarrow> ep B A"
|
|
31 |
and "\<And>A B C. ep A B \<Longrightarrow> ep B C \<Longrightarrow> ep A C"
|
|
32 |
and "\<And>A B. pl A B \<Longrightarrow> el B B"
|
|
33 |
and "\<And>A B. el A B \<Longrightarrow> el B A"
|
|
34 |
and "\<And>A B C. el A B \<Longrightarrow> el B C \<Longrightarrow> el A C"
|
|
35 |
and "\<And>A B C. ep A B \<Longrightarrow> pl B C \<Longrightarrow> pl A C"
|
|
36 |
and "\<And>A B C. pl A B \<Longrightarrow> el B C \<Longrightarrow> pl A C"
|
|
37 |
and "\<And>A B C D E F G H I J K L M N O P Q.
|
|
38 |
col A B C D \<Longrightarrow> col E F G H \<Longrightarrow> col B G I J \<Longrightarrow> col C F I K \<Longrightarrow>
|
|
39 |
col B E L M \<Longrightarrow> col A F L N \<Longrightarrow> col C E O P \<Longrightarrow> col A G O Q \<Longrightarrow>
|
|
40 |
(\<exists> R. col I L O R) \<or> pl A H \<or> pl B H \<or> pl C H \<or> pl E D \<or> pl F D \<or> pl G D"
|
|
41 |
and "\<And>A B C D. pl A B \<Longrightarrow> pl A C \<Longrightarrow> pl D B \<Longrightarrow> pl D C \<Longrightarrow> ep A D \<or> el B C"
|
|
42 |
and "\<And>A B. ep A A \<Longrightarrow> ep B B \<Longrightarrow> \<exists>C. pl A C \<and> pl B C"
|
|
43 |
shows goal using assms
|
|
44 |
by coherent
|
|
45 |
|
|
46 |
lemma p2p1:
|
|
47 |
assumes "col a b c l \<and> col d e f m"
|
|
48 |
and "col b f g n \<and> col c e g o"
|
|
49 |
and "col b d h p \<and> col a e h q"
|
|
50 |
and "col c d i r \<and> col a f i s"
|
|
51 |
and "pl a m \<Longrightarrow> goal"
|
|
52 |
and "pl b m \<Longrightarrow> goal"
|
|
53 |
and "pl c m \<Longrightarrow> goal"
|
|
54 |
and "pl d l \<Longrightarrow> goal"
|
|
55 |
and "pl e l \<Longrightarrow> goal"
|
|
56 |
and "pl f l \<Longrightarrow> goal"
|
|
57 |
and "\<And>A. pl g A \<Longrightarrow> pl h A \<Longrightarrow> pl i A \<Longrightarrow> goal"
|
|
58 |
and "\<And>A B C D. col A B C D \<Longrightarrow> pl A D"
|
|
59 |
and "\<And>A B C D. col A B C D \<Longrightarrow> pl B D"
|
|
60 |
and "\<And>A B C D. col A B C D \<Longrightarrow> pl C D"
|
|
61 |
and "\<And>A B. pl A B \<Longrightarrow> ep A A"
|
|
62 |
and "\<And>A B. ep A B \<Longrightarrow> ep B A"
|
|
63 |
and "\<And>A B C. ep A B \<Longrightarrow> ep B C \<Longrightarrow> ep A C"
|
|
64 |
and "\<And>A B. pl A B \<Longrightarrow> el B B"
|
|
65 |
and "\<And>A B. el A B \<Longrightarrow> el B A"
|
|
66 |
and "\<And>A B C. el A B \<Longrightarrow> el B C \<Longrightarrow> el A C"
|
|
67 |
and "\<And>A B C. ep A B \<Longrightarrow> pl B C \<Longrightarrow> pl A C"
|
|
68 |
and "\<And>A B C. pl A B \<Longrightarrow> el B C \<Longrightarrow> pl A C"
|
|
69 |
and "\<And>A B C D E F G H I J K L M N O P Q.
|
|
70 |
col A B C J \<Longrightarrow> col D E F K \<Longrightarrow> col B F G L \<Longrightarrow> col C E G M \<Longrightarrow>
|
|
71 |
col B D H N \<Longrightarrow> col A E H O \<Longrightarrow> col C D I P \<Longrightarrow> col A F I Q \<Longrightarrow>
|
|
72 |
(\<exists> R. col G H I R) \<or> el L M \<or> el N O \<or> el P Q"
|
|
73 |
and "\<And>A B C D. pl C A \<Longrightarrow> pl C B \<Longrightarrow> pl D A \<Longrightarrow> pl D B \<Longrightarrow> ep C D \<or> el A B"
|
|
74 |
and "\<And>A B C. ep A A \<Longrightarrow> ep B B \<Longrightarrow> \<exists>C. pl A C \<and> pl B C"
|
|
75 |
shows goal using assms
|
|
76 |
by coherent
|
|
77 |
|
|
78 |
|
|
79 |
subsection \<open>Preservation of the Diamond Property under reflexive closure\<close>
|
|
80 |
|
|
81 |
lemma diamond:
|
|
82 |
assumes "reflexive_rewrite a b" "reflexive_rewrite a c"
|
|
83 |
and "\<And>A. reflexive_rewrite b A \<Longrightarrow> reflexive_rewrite c A \<Longrightarrow> goal"
|
|
84 |
and "\<And>A. equalish A A"
|
|
85 |
and "\<And>A B. equalish A B \<Longrightarrow> equalish B A"
|
|
86 |
and "\<And>A B C. equalish A B \<Longrightarrow> reflexive_rewrite B C \<Longrightarrow> reflexive_rewrite A C"
|
|
87 |
and "\<And>A B. equalish A B \<Longrightarrow> reflexive_rewrite A B"
|
|
88 |
and "\<And>A B. rewrite A B \<Longrightarrow> reflexive_rewrite A B"
|
|
89 |
and "\<And>A B. reflexive_rewrite A B \<Longrightarrow> equalish A B \<or> rewrite A B"
|
|
90 |
and "\<And>A B C. rewrite A B \<Longrightarrow> rewrite A C \<Longrightarrow> \<exists>D. rewrite B D \<and> rewrite C D"
|
|
91 |
shows goal using assms
|
|
92 |
by coherent
|
|
93 |
|
|
94 |
end
|