src/FOL/ex/Propositional_Int.thy
author immler
Sun, 27 Oct 2019 21:51:14 -0400
changeset 71035 6fe5a0e1fa8e
parent 69593 3dda49e08b9d
permissions -rw-r--r--
moved theory Interval from the AFP
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
     1
(*  Title:      FOL/ex/Propositional_Int.thy
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
     3
    Copyright   1991  University of Cambridge
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
     4
*)
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
     5
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     6
section \<open>First-Order Logic: propositional examples (intuitionistic version)\<close>
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
     7
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
     8
theory Propositional_Int
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
     9
imports IFOL
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    10
begin
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    11
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61489
diff changeset
    12
text \<open>commutative laws of \<open>\<and>\<close> and \<open>\<or>\<close>\<close>
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    13
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    14
lemma \<open>P \<and> Q \<longrightarrow> Q \<and> P\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    15
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    16
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    17
lemma \<open>P \<or> Q \<longrightarrow> Q \<or> P\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    18
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    19
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    20
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61489
diff changeset
    21
text \<open>associative laws of \<open>\<and>\<close> and \<open>\<or>\<close>\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    22
lemma \<open>(P \<and> Q) \<and> R \<longrightarrow> P \<and> (Q \<and> R)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    23
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    24
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    25
lemma \<open>(P \<or> Q) \<or> R \<longrightarrow> P \<or> (Q \<or> R)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    26
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    27
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    28
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61489
diff changeset
    29
text \<open>distributive laws of \<open>\<and>\<close> and \<open>\<or>\<close>\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    30
lemma \<open>(P \<and> Q) \<or> R \<longrightarrow> (P \<or> R) \<and> (Q \<or> R)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    31
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    32
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    33
lemma \<open>(P \<or> R) \<and> (Q \<or> R) \<longrightarrow> (P \<and> Q) \<or> R\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    34
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    35
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    36
lemma \<open>(P \<or> Q) \<and> R \<longrightarrow> (P \<and> R) \<or> (Q \<and> R)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    37
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    38
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    39
lemma \<open>(P \<and> R) \<or> (Q \<and> R) \<longrightarrow> (P \<or> Q) \<and> R\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    40
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    41
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    42
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    43
text \<open>Laws involving implication\<close>
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    44
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    45
lemma \<open>(P \<longrightarrow> R) \<and> (Q \<longrightarrow> R) \<longleftrightarrow> (P \<or> Q \<longrightarrow> R)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    46
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    47
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    48
lemma \<open>(P \<and> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    49
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    50
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    51
lemma \<open>((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> ((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> (P \<and> Q \<longrightarrow> R) \<longrightarrow> R\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    52
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    53
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    54
lemma \<open>\<not> (P \<longrightarrow> R) \<longrightarrow> \<not> (Q \<longrightarrow> R) \<longrightarrow> \<not> (P \<and> Q \<longrightarrow> R)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    55
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    56
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    57
lemma \<open>(P \<longrightarrow> Q \<and> R) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    58
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    59
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    60
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    61
text \<open>Propositions-as-types\<close>
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    62
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61489
diff changeset
    63
\<comment> \<open>The combinator K\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    64
lemma \<open>P \<longrightarrow> (Q \<longrightarrow> P)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    65
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    66
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61489
diff changeset
    67
\<comment> \<open>The combinator S\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    68
lemma \<open>(P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q) \<longrightarrow> (P \<longrightarrow> R)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    69
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    70
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    71
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61489
diff changeset
    72
\<comment> \<open>Converse is classical\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    73
lemma \<open>(P \<longrightarrow> Q) \<or> (P \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q \<or> R)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    74
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    75
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    76
lemma \<open>(P \<longrightarrow> Q) \<longrightarrow> (\<not> Q \<longrightarrow> \<not> P)\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    77
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    78
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    79
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    80
text \<open>Schwichtenberg's examples (via T. Nipkow)\<close>
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    81
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    82
lemma stab_imp: \<open>(((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> (((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> P \<longrightarrow> Q\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    83
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    84
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    85
lemma stab_to_peirce:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    86
  \<open>(((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> (((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> Q)
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    87
    \<longrightarrow> ((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    88
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    89
61489
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
    90
lemma peirce_imp1:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    91
  \<open>(((Q \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> Q)
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    92
    \<longrightarrow> (((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> P \<longrightarrow> Q) \<longrightarrow> P \<longrightarrow> Q\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    93
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
    94
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    95
lemma peirce_imp2: \<open>(((P \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> ((P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> P) \<longrightarrow> P\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    96
  by (tactic "IntPr.fast_tac \<^context> 1")
61489
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
    97
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
    98
lemma mints: \<open>((((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> Q) \<longrightarrow> Q\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
    99
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
   100
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
   101
lemma mints_solovev: \<open>(P \<longrightarrow> (Q \<longrightarrow> R) \<longrightarrow> Q) \<longrightarrow> ((P \<longrightarrow> Q) \<longrightarrow> R) \<longrightarrow> R\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
   102
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
   103
61489
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
   104
lemma tatsuta:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
   105
  \<open>(((P7 \<longrightarrow> P1) \<longrightarrow> P10) \<longrightarrow> P4 \<longrightarrow> P5)
61489
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
   106
  \<longrightarrow> (((P8 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P3 \<longrightarrow> P10)
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
   107
  \<longrightarrow> (P1 \<longrightarrow> P8) \<longrightarrow> P6 \<longrightarrow> P7
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
   108
  \<longrightarrow> (((P3 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P4)
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
   109
  \<longrightarrow> (P1 \<longrightarrow> P3) \<longrightarrow> (((P6 \<longrightarrow> P1) \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P5\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
   110
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
   111
61489
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
   112
lemma tatsuta1:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
   113
  \<open>(((P8 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P3 \<longrightarrow> P10)
61489
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
   114
  \<longrightarrow> (((P3 \<longrightarrow> P2) \<longrightarrow> P9) \<longrightarrow> P4)
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
   115
  \<longrightarrow> (((P6 \<longrightarrow> P1) \<longrightarrow> P2) \<longrightarrow> P9)
b8d375aee0df more symbols;
wenzelm
parents: 60770
diff changeset
   116
  \<longrightarrow> (((P7 \<longrightarrow> P1) \<longrightarrow> P10) \<longrightarrow> P4 \<longrightarrow> P5)
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 62020
diff changeset
   117
  \<longrightarrow> (P1 \<longrightarrow> P3) \<longrightarrow> (P1 \<longrightarrow> P8) \<longrightarrow> P6 \<longrightarrow> P7 \<longrightarrow> P5\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
   118
  by (tactic "IntPr.fast_tac \<^context> 1")
23914
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
   119
3e0424305fa4 turned ex/prop.ML, ex/quant.ML into proper theories;
wenzelm
parents:
diff changeset
   120
end