src/FOLP/ex/Propositional_Int.thy
changeset 69593 3dda49e08b9d
parent 61337 4645502c3c64
equal deleted inserted replaced
69592:a80d8ec6c998 69593:3dda49e08b9d
    10 begin
    10 begin
    11 
    11 
    12 
    12 
    13 text "commutative laws of & and | "
    13 text "commutative laws of & and | "
    14 schematic_goal "?p : P & Q  -->  Q & P"
    14 schematic_goal "?p : P & Q  -->  Q & P"
    15   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    15   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    16 
    16 
    17 schematic_goal "?p : P | Q  -->  Q | P"
    17 schematic_goal "?p : P | Q  -->  Q | P"
    18   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    18   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    19 
    19 
    20 
    20 
    21 text "associative laws of & and | "
    21 text "associative laws of & and | "
    22 schematic_goal "?p : (P & Q) & R  -->  P & (Q & R)"
    22 schematic_goal "?p : (P & Q) & R  -->  P & (Q & R)"
    23   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    23   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    24 
    24 
    25 schematic_goal "?p : (P | Q) | R  -->  P | (Q | R)"
    25 schematic_goal "?p : (P | Q) | R  -->  P | (Q | R)"
    26   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    26   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    27 
    27 
    28 
    28 
    29 text "distributive laws of & and | "
    29 text "distributive laws of & and | "
    30 schematic_goal "?p : (P & Q) | R  --> (P | R) & (Q | R)"
    30 schematic_goal "?p : (P & Q) | R  --> (P | R) & (Q | R)"
    31   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    31   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    32 
    32 
    33 schematic_goal "?p : (P | R) & (Q | R)  --> (P & Q) | R"
    33 schematic_goal "?p : (P | R) & (Q | R)  --> (P & Q) | R"
    34   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    34   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    35 
    35 
    36 schematic_goal "?p : (P | Q) & R  --> (P & R) | (Q & R)"
    36 schematic_goal "?p : (P | Q) & R  --> (P & R) | (Q & R)"
    37   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    37   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    38 
    38 
    39 
    39 
    40 schematic_goal "?p : (P & R) | (Q & R)  --> (P | Q) & R"
    40 schematic_goal "?p : (P & R) | (Q & R)  --> (P | Q) & R"
    41   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    41   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    42 
    42 
    43 
    43 
    44 text "Laws involving implication"
    44 text "Laws involving implication"
    45 
    45 
    46 schematic_goal "?p : (P-->R) & (Q-->R) <-> (P|Q --> R)"
    46 schematic_goal "?p : (P-->R) & (Q-->R) <-> (P|Q --> R)"
    47   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    47   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    48 
    48 
    49 schematic_goal "?p : (P & Q --> R) <-> (P--> (Q-->R))"
    49 schematic_goal "?p : (P & Q --> R) <-> (P--> (Q-->R))"
    50   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    50   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    51 
    51 
    52 schematic_goal "?p : ((P-->R)-->R) --> ((Q-->R)-->R) --> (P&Q-->R) --> R"
    52 schematic_goal "?p : ((P-->R)-->R) --> ((Q-->R)-->R) --> (P&Q-->R) --> R"
    53   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    53   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    54 
    54 
    55 schematic_goal "?p : ~(P-->R) --> ~(Q-->R) --> ~(P&Q-->R)"
    55 schematic_goal "?p : ~(P-->R) --> ~(Q-->R) --> ~(P&Q-->R)"
    56   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    56   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    57 
    57 
    58 schematic_goal "?p : (P --> Q & R) <-> (P-->Q)  &  (P-->R)"
    58 schematic_goal "?p : (P --> Q & R) <-> (P-->Q)  &  (P-->R)"
    59   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    59   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    60 
    60 
    61 
    61 
    62 text "Propositions-as-types"
    62 text "Propositions-as-types"
    63 
    63 
    64 (*The combinator K*)
    64 (*The combinator K*)
    65 schematic_goal "?p : P --> (Q --> P)"
    65 schematic_goal "?p : P --> (Q --> P)"
    66   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    66   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    67 
    67 
    68 (*The combinator S*)
    68 (*The combinator S*)
    69 schematic_goal "?p : (P-->Q-->R)  --> (P-->Q) --> (P-->R)"
    69 schematic_goal "?p : (P-->Q-->R)  --> (P-->Q) --> (P-->R)"
    70   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    70   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    71 
    71 
    72 
    72 
    73 (*Converse is classical*)
    73 (*Converse is classical*)
    74 schematic_goal "?p : (P-->Q) | (P-->R)  -->  (P --> Q | R)"
    74 schematic_goal "?p : (P-->Q) | (P-->R)  -->  (P --> Q | R)"
    75   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    75   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    76 
    76 
    77 schematic_goal "?p : (P-->Q)  -->  (~Q --> ~P)"
    77 schematic_goal "?p : (P-->Q)  -->  (~Q --> ~P)"
    78   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    78   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    79 
    79 
    80 
    80 
    81 text "Schwichtenberg's examples (via T. Nipkow)"
    81 text "Schwichtenberg's examples (via T. Nipkow)"
    82 
    82 
    83 schematic_goal stab_imp: "?p : (((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
    83 schematic_goal stab_imp: "?p : (((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
    84   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    84   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    85 
    85 
    86 schematic_goal stab_to_peirce: "?p : (((P --> R) --> R) --> P) --> (((Q --> R) --> R) --> Q)  
    86 schematic_goal stab_to_peirce: "?p : (((P --> R) --> R) --> P) --> (((Q --> R) --> R) --> Q)  
    87               --> ((P --> Q) --> P) --> P"
    87               --> ((P --> Q) --> P) --> P"
    88   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    88   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    89 
    89 
    90 schematic_goal peirce_imp1: "?p : (((Q --> R) --> Q) --> Q)  
    90 schematic_goal peirce_imp1: "?p : (((Q --> R) --> Q) --> Q)  
    91                --> (((P --> Q) --> R) --> P --> Q) --> P --> Q"
    91                --> (((P --> Q) --> R) --> P --> Q) --> P --> Q"
    92   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    92   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    93   
    93   
    94 schematic_goal peirce_imp2: "?p : (((P --> R) --> P) --> P) --> ((P --> Q --> R) --> P) --> P"
    94 schematic_goal peirce_imp2: "?p : (((P --> R) --> P) --> P) --> ((P --> Q --> R) --> P) --> P"
    95   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    95   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    96 
    96 
    97 schematic_goal mints: "?p : ((((P --> Q) --> P) --> P) --> Q) --> Q"
    97 schematic_goal mints: "?p : ((((P --> Q) --> P) --> P) --> Q) --> Q"
    98   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
    98   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
    99 
    99 
   100 schematic_goal mints_solovev: "?p : (P --> (Q --> R) --> Q) --> ((P --> Q) --> R) --> R"
   100 schematic_goal mints_solovev: "?p : (P --> (Q --> R) --> Q) --> ((P --> Q) --> R) --> R"
   101   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
   101   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
   102 
   102 
   103 schematic_goal tatsuta: "?p : (((P7 --> P1) --> P10) --> P4 --> P5)  
   103 schematic_goal tatsuta: "?p : (((P7 --> P1) --> P10) --> P4 --> P5)  
   104           --> (((P8 --> P2) --> P9) --> P3 --> P10)  
   104           --> (((P8 --> P2) --> P9) --> P3 --> P10)  
   105           --> (P1 --> P8) --> P6 --> P7  
   105           --> (P1 --> P8) --> P6 --> P7  
   106           --> (((P3 --> P2) --> P9) --> P4)  
   106           --> (((P3 --> P2) --> P9) --> P4)  
   107           --> (P1 --> P3) --> (((P6 --> P1) --> P2) --> P9) --> P5"
   107           --> (P1 --> P3) --> (((P6 --> P1) --> P2) --> P9) --> P5"
   108   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
   108   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
   109 
   109 
   110 schematic_goal tatsuta1: "?p : (((P8 --> P2) --> P9) --> P3 --> P10)  
   110 schematic_goal tatsuta1: "?p : (((P8 --> P2) --> P9) --> P3 --> P10)  
   111      --> (((P3 --> P2) --> P9) --> P4)  
   111      --> (((P3 --> P2) --> P9) --> P4)  
   112      --> (((P6 --> P1) --> P2) --> P9)  
   112      --> (((P6 --> P1) --> P2) --> P9)  
   113      --> (((P7 --> P1) --> P10) --> P4 --> P5)  
   113      --> (((P7 --> P1) --> P10) --> P4 --> P5)  
   114      --> (P1 --> P3) --> (P1 --> P8) --> P6 --> P7 --> P5"
   114      --> (P1 --> P3) --> (P1 --> P8) --> P6 --> P7 --> P5"
   115   by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
   115   by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)
   116 
   116 
   117 end
   117 end