src/FOL/ex/Propositional_Int.thy
changeset 60770 240563fbf41d
parent 58889 5b7a9633cfa8
child 61489 b8d375aee0df
--- a/src/FOL/ex/Propositional_Int.thy	Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Propositional_Int.thy	Thu Jul 23 14:25:05 2015 +0200
@@ -3,13 +3,13 @@
     Copyright   1991  University of Cambridge
 *)
 
-section {* First-Order Logic: propositional examples (intuitionistic version) *}
+section \<open>First-Order Logic: propositional examples (intuitionistic version)\<close>
 
 theory Propositional_Int
 imports IFOL
 begin
 
-text {* commutative laws of @{text "&"} and @{text "|"} *}
+text \<open>commutative laws of @{text "&"} and @{text "|"}\<close>
 
 lemma "P & Q  -->  Q & P"
   by (tactic "IntPr.fast_tac @{context} 1")
@@ -18,7 +18,7 @@
   by (tactic "IntPr.fast_tac @{context} 1")
 
 
-text {* associative laws of @{text "&"} and @{text "|"} *}
+text \<open>associative laws of @{text "&"} and @{text "|"}\<close>
 lemma "(P & Q) & R  -->  P & (Q & R)"
   by (tactic "IntPr.fast_tac @{context} 1")
 
@@ -26,7 +26,7 @@
   by (tactic "IntPr.fast_tac @{context} 1")
 
 
-text {* distributive laws of @{text "&"} and @{text "|"} *}
+text \<open>distributive laws of @{text "&"} and @{text "|"}\<close>
 lemma "(P & Q) | R  --> (P | R) & (Q | R)"
   by (tactic "IntPr.fast_tac @{context} 1")
 
@@ -40,7 +40,7 @@
   by (tactic "IntPr.fast_tac @{context} 1")
 
 
-text {* Laws involving implication *}
+text \<open>Laws involving implication\<close>
 
 lemma "(P-->R) & (Q-->R) <-> (P|Q --> R)"
   by (tactic "IntPr.fast_tac @{context} 1")
@@ -58,18 +58,18 @@
   by (tactic "IntPr.fast_tac @{context} 1")
 
 
-text {* Propositions-as-types *}
+text \<open>Propositions-as-types\<close>
 
--- {* The combinator K *}
+-- \<open>The combinator K\<close>
 lemma "P --> (Q --> P)"
   by (tactic "IntPr.fast_tac @{context} 1")
 
--- {* The combinator S *}
+-- \<open>The combinator S\<close>
 lemma "(P-->Q-->R)  --> (P-->Q) --> (P-->R)"
   by (tactic "IntPr.fast_tac @{context} 1")
 
 
--- {* Converse is classical *}
+-- \<open>Converse is classical\<close>
 lemma "(P-->Q) | (P-->R)  -->  (P --> Q | R)"
   by (tactic "IntPr.fast_tac @{context} 1")
 
@@ -77,7 +77,7 @@
   by (tactic "IntPr.fast_tac @{context} 1")
 
 
-text {* Schwichtenberg's examples (via T. Nipkow) *}
+text \<open>Schwichtenberg's examples (via T. Nipkow)\<close>
 
 lemma stab_imp: "(((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
   by (tactic "IntPr.fast_tac @{context} 1")