src/Cube/Example.thy
changeset 62020 5d208fd2507d
parent 61390 a705f42b244d
--- a/src/Cube/Example.thy	Thu Dec 31 21:46:31 2015 +0100
+++ b/src/Cube/Example.thy	Fri Jan 01 10:49:00 2016 +0100
@@ -149,7 +149,7 @@
 
 schematic_goal (in LP2) "A:* P:A\<rightarrow>A\<rightarrow>* \<turnstile>
     ?p: (\<Prod>a:A. \<Prod>b:A. P\<cdot>a\<cdot>b\<rightarrow>P\<cdot>b\<cdot>a\<rightarrow>\<Prod>P:*.P) \<rightarrow> \<Prod>a:A. P\<cdot>a\<cdot>a\<rightarrow>\<Prod>P:*.P"
-  -- \<open>Antisymmetry implies irreflexivity:\<close>
+  \<comment> \<open>Antisymmetry implies irreflexivity:\<close>
   apply (strip_asms rules)
   apply (rule lam_ss)
     apply (depth_solve1 rules)
@@ -206,7 +206,7 @@
 
 schematic_goal (in LP2)
   "A:* a:A b:A \<turnstile> ?p: (\<Prod>P:A\<rightarrow>*.P\<cdot>a\<rightarrow>P\<cdot>b) \<rightarrow> (\<Prod>P:A\<rightarrow>*.P\<cdot>b\<rightarrow>P\<cdot>a)"
-  -- \<open>Symmetry of Leibnitz equality\<close>
+  \<comment> \<open>Symmetry of Leibnitz equality\<close>
   apply (strip_asms rules)
   apply (rule lam_ss)
     apply (depth_solve1 rules)