diff -r cf7f3465eaf1 -r 240563fbf41d src/FOL/ex/Propositional_Int.thy --- 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 \First-Order Logic: propositional examples (intuitionistic version)\ theory Propositional_Int imports IFOL begin -text {* commutative laws of @{text "&"} and @{text "|"} *} +text \commutative laws of @{text "&"} and @{text "|"}\ 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 \associative laws of @{text "&"} and @{text "|"}\ 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 \distributive laws of @{text "&"} and @{text "|"}\ 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 \Laws involving implication\ 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 \Propositions-as-types\ --- {* The combinator K *} +-- \The combinator K\ lemma "P --> (Q --> P)" by (tactic "IntPr.fast_tac @{context} 1") --- {* The combinator S *} +-- \The combinator S\ lemma "(P-->Q-->R) --> (P-->Q) --> (P-->R)" by (tactic "IntPr.fast_tac @{context} 1") --- {* Converse is classical *} +-- \Converse is classical\ 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 \Schwichtenberg's examples (via T. Nipkow)\ lemma stab_imp: "(((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q" by (tactic "IntPr.fast_tac @{context} 1")