src/HOL/ex/Coherent.thy
author wenzelm
Tue, 05 Nov 2019 14:16:16 +0100
changeset 71046 b8aeeedf7e68
parent 63054 1b237d147cc4
permissions -rw-r--r--
support for Linux user management;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32734
06c13b2e562e misc tuning and modernization;
wenzelm
parents: 28323
diff changeset
     1
(*  Title:      HOL/ex/Coherent.thy
28323
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
32734
06c13b2e562e misc tuning and modernization;
wenzelm
parents: 28323
diff changeset
     3
    Author:     Marc Bezem, Institutt for Informatikk, Universitetet i Bergen 
28323
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
     4
*)
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
     5
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     6
section \<open>Coherent Logic Problems\<close>
28323
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
     7
32734
06c13b2e562e misc tuning and modernization;
wenzelm
parents: 28323
diff changeset
     8
theory Coherent
06c13b2e562e misc tuning and modernization;
wenzelm
parents: 28323
diff changeset
     9
imports Main
06c13b2e562e misc tuning and modernization;
wenzelm
parents: 28323
diff changeset
    10
begin
28323
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    11
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    12
subsection \<open>Equivalence of two versions of Pappus' Axiom\<close>
28323
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    13
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    14
no_notation
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    15
  comp (infixl "o" 55) and
47433
07f4bf913230 renamed "rel_comp" to "relcomp" (to be consistent with, e.g., "relpow")
griff
parents: 32734
diff changeset
    16
  relcomp (infixr "O" 75)
28323
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    17
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    18
lemma p1p2:
63054
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    19
  assumes "col a b c l \<and> col d e f m"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    20
    and "col b f g n \<and> col c e g o"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    21
    and "col b d h p \<and> col a e h q"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    22
    and "col c d i r \<and> col a f i s"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    23
    and "el n o \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    24
    and "el p q \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    25
    and "el s r \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    26
    and "\<And>A. el A A \<Longrightarrow> pl g A \<Longrightarrow> pl h A \<Longrightarrow> pl i A \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    27
    and "\<And>A B C D. col A B C D \<Longrightarrow> pl A D"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    28
    and "\<And>A B C D. col A B C D \<Longrightarrow> pl B D"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    29
    and "\<And>A B C D. col A B C D \<Longrightarrow> pl C D"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    30
    and "\<And>A B. pl A B \<Longrightarrow> ep A A"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    31
    and "\<And>A B. ep A B \<Longrightarrow> ep B A"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    32
    and "\<And>A B C. ep A B \<Longrightarrow> ep B C \<Longrightarrow> ep A C"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    33
    and "\<And>A B. pl A B \<Longrightarrow> el B B"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    34
    and "\<And>A B. el A B \<Longrightarrow> el B A"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    35
    and "\<And>A B C. el A B \<Longrightarrow> el B C \<Longrightarrow> el A C"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    36
    and "\<And>A B C. ep A B \<Longrightarrow> pl B C \<Longrightarrow> pl A C"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    37
    and "\<And>A B C. pl A B \<Longrightarrow> el B C \<Longrightarrow> pl A C"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    38
    and "\<And>A B C D E F G H I J K L M N O P Q.
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    39
           col A B C D \<Longrightarrow> col E F G H \<Longrightarrow> col B G I J \<Longrightarrow> col C F I K \<Longrightarrow>
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    40
           col B E L M \<Longrightarrow> col A F L N \<Longrightarrow> col C E O P \<Longrightarrow> col A G O Q \<Longrightarrow>
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    41
           (\<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"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    42
    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"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    43
    and "\<And>A B. ep A A \<Longrightarrow> ep B B \<Longrightarrow> \<exists>C. pl A C \<and> pl B C"
28323
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    44
  shows goal using assms
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    45
  by coherent
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    46
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    47
lemma p2p1:
63054
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    48
  assumes "col a b c l \<and> col d e f m"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    49
    and "col b f g n \<and> col c e g o"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    50
    and "col b d h p \<and> col a e h q"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    51
    and "col c d i r \<and> col a f i s"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    52
    and "pl a m \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    53
    and "pl b m \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    54
    and "pl c m \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    55
    and "pl d l \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    56
    and "pl e l \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    57
    and "pl f l \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    58
    and "\<And>A. pl g A \<Longrightarrow> pl h A \<Longrightarrow> pl i A \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    59
    and "\<And>A B C D. col A B C D \<Longrightarrow> pl A D"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    60
    and "\<And>A B C D. col A B C D \<Longrightarrow> pl B D"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    61
    and "\<And>A B C D. col A B C D \<Longrightarrow> pl C D"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    62
    and "\<And>A B. pl A B \<Longrightarrow> ep A A"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    63
    and "\<And>A B. ep A B \<Longrightarrow> ep B A"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    64
    and "\<And>A B C. ep A B \<Longrightarrow> ep B C \<Longrightarrow> ep A C"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    65
    and "\<And>A B. pl A B \<Longrightarrow> el B B"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    66
    and "\<And>A B. el A B \<Longrightarrow> el B A"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    67
    and "\<And>A B C. el A B \<Longrightarrow> el B C \<Longrightarrow> el A C"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    68
    and "\<And>A B C. ep A B \<Longrightarrow> pl B C \<Longrightarrow> pl A C"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    69
    and "\<And>A B C. pl A B \<Longrightarrow> el B C \<Longrightarrow> pl A C"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    70
    and "\<And>A B C D E F G H I J K L M N O P Q.
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    71
           col A B C J \<Longrightarrow> col D E F K \<Longrightarrow> col B F G L \<Longrightarrow> col C E G M \<Longrightarrow>
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    72
           col B D H N \<Longrightarrow> col A E H O \<Longrightarrow> col C D I P \<Longrightarrow> col A F I Q \<Longrightarrow>
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    73
           (\<exists> R. col G H I R) \<or> el L M \<or> el N O \<or> el P Q"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    74
    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"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    75
    and "\<And>A B C. ep A A \<Longrightarrow> ep B B \<Longrightarrow> \<exists>C. pl A C \<and> pl B C"
28323
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    76
  shows goal using assms
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    77
  by coherent
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    78
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    79
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    80
subsection \<open>Preservation of the Diamond Property under reflexive closure\<close>
28323
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    81
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    82
lemma diamond:
63054
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    83
  assumes "reflexive_rewrite a b" "reflexive_rewrite a c"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    84
    and "\<And>A. reflexive_rewrite b A \<Longrightarrow> reflexive_rewrite c A \<Longrightarrow> goal"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    85
    and "\<And>A. equalish A A" 
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    86
    and "\<And>A B. equalish A B \<Longrightarrow> equalish B A"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    87
    and "\<And>A B C. equalish A B \<Longrightarrow> reflexive_rewrite B C \<Longrightarrow> reflexive_rewrite A C"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    88
    and "\<And>A B. equalish A B \<Longrightarrow> reflexive_rewrite A B"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    89
    and "\<And>A B. rewrite A B \<Longrightarrow> reflexive_rewrite A B"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    90
    and "\<And>A B. reflexive_rewrite A B \<Longrightarrow> equalish A B \<or> rewrite A B"
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61343
diff changeset
    91
    and "\<And>A B C. rewrite A B \<Longrightarrow> rewrite A C \<Longrightarrow> \<exists>D. rewrite B D \<and> rewrite C D"
28323
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    92
  shows goal using assms
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    93
  by coherent
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    94
8f12f7275637 Examples for coherent logic prover.
berghofe
parents:
diff changeset
    95
end