src/Cube/Example.thy
changeset 58617 4f169d2cf6f3
parent 45242 401f91ed8a93
child 58889 5b7a9633cfa8
--- a/src/Cube/Example.thy	Tue Oct 07 21:11:18 2014 +0200
+++ b/src/Cube/Example.thy	Tue Oct 07 21:28:24 2014 +0200
@@ -1,34 +1,29 @@
-header {* Lambda Cube Examples *}
+header \<open>Lambda Cube Examples\<close>
 
 theory Example
 imports Cube
 begin
 
-text {*
-  Examples taken from:
+text \<open>Examples taken from:
 
   H. Barendregt. Introduction to Generalised Type Systems.
-  J. Functional Programming.
-*}
+  J. Functional Programming.\<close>
 
-method_setup depth_solve = {*
-  Attrib.thms >> (fn thms => K (METHOD (fn facts =>
-    (DEPTH_SOLVE (HEADGOAL (ares_tac (facts @ thms)))))))
-*}
+method_setup depth_solve =
+  \<open>Attrib.thms >> (fn thms => K (METHOD (fn facts =>
+    (DEPTH_SOLVE (HEADGOAL (ares_tac (facts @ thms)))))))\<close>
 
-method_setup depth_solve1 = {*
-  Attrib.thms >> (fn thms => K (METHOD (fn facts =>
-    (DEPTH_SOLVE_1 (HEADGOAL (ares_tac (facts @ thms)))))))
-*}
+method_setup depth_solve1 =
+  \<open>Attrib.thms >> (fn thms => K (METHOD (fn facts =>
+    (DEPTH_SOLVE_1 (HEADGOAL (ares_tac (facts @ thms)))))))\<close>
 
-method_setup strip_asms =  {*
-  Attrib.thms >> (fn thms => K (METHOD (fn facts =>
+method_setup strip_asms =
+  \<open>Attrib.thms >> (fn thms => K (METHOD (fn facts =>
     REPEAT (resolve_tac [@{thm strip_b}, @{thm strip_s}] 1 THEN
-    DEPTH_SOLVE_1 (ares_tac (facts @ thms) 1)))))
-*}
+    DEPTH_SOLVE_1 (ares_tac (facts @ thms) 1)))))\<close>
 
 
-subsection {* Simple types *}
+subsection \<open>Simple types\<close>
 
 schematic_lemma "A:* \<turnstile> A\<rightarrow>A : ?T"
   by (depth_solve rules)
@@ -49,7 +44,7 @@
   by (depth_solve rules)
 
 
-subsection {* Second-order types *}
+subsection \<open>Second-order types\<close>
 
 schematic_lemma (in L2) "\<turnstile> \<Lambda> A:*. \<Lambda> a:A. a : ?T"
   by (depth_solve rules)
@@ -64,7 +59,7 @@
   by (depth_solve rules)
 
 
-subsection {* Weakly higher-order propositional logic *}
+subsection \<open>Weakly higher-order propositional logic\<close>
 
 schematic_lemma (in Lomega) "\<turnstile> \<Lambda> A:*.A\<rightarrow>A : ?T"
   by (depth_solve rules)
@@ -82,7 +77,7 @@
   by (depth_solve rules)
 
 
-subsection {* LP *}
+subsection \<open>LP\<close>
 
 schematic_lemma (in LP) "A:* \<turnstile> A \<rightarrow> * : ?T"
   by (depth_solve rules)
@@ -110,7 +105,7 @@
   by (depth_solve rules)
 
 
-subsection {* Omega-order types *}
+subsection \<open>Omega-order types\<close>
 
 schematic_lemma (in L2) "A:* B:* \<turnstile> \<Pi> C:*.(A\<rightarrow>B\<rightarrow>C)\<rightarrow>C : ?T"
   by (depth_solve rules)
@@ -143,7 +138,7 @@
   done
 
 
-subsection {* Second-order Predicate Logic *}
+subsection \<open>Second-order Predicate Logic\<close>
 
 schematic_lemma (in LP2) "A:* P:A\<rightarrow>* \<turnstile> \<Lambda> a:A. P^a\<rightarrow>(\<Pi> A:*.A) : ?T"
   by (depth_solve rules)
@@ -154,7 +149,7 @@
 
 schematic_lemma (in LP2) "A:* P:A\<rightarrow>A\<rightarrow>* \<turnstile>
     ?p: (\<Pi> a:A. \<Pi> b:A. P^a^b\<rightarrow>P^b^a\<rightarrow>\<Pi> P:*.P) \<rightarrow> \<Pi> a:A. P^a^a\<rightarrow>\<Pi> P:*.P"
-  -- {* Antisymmetry implies irreflexivity: *}
+  -- \<open>Antisymmetry implies irreflexivity:\<close>
   apply (strip_asms rules)
   apply (rule lam_ss)
     apply (depth_solve1 rules)
@@ -172,7 +167,7 @@
   done
 
 
-subsection {* LPomega *}
+subsection \<open>LPomega\<close>
 
 schematic_lemma (in LPomega) "A:* \<turnstile> \<Lambda> P:A\<rightarrow>A\<rightarrow>*.\<Lambda> a:A. P^a^a : ?T"
   by (depth_solve rules)
@@ -181,7 +176,7 @@
   by (depth_solve rules)
 
 
-subsection {* Constructions *}
+subsection \<open>Constructions\<close>
 
 schematic_lemma (in CC) "\<turnstile> \<Lambda> A:*.\<Lambda> P:A\<rightarrow>*.\<Lambda> a:A. P^a\<rightarrow>\<Pi> P:*.P: ?T"
   by (depth_solve rules)
@@ -199,7 +194,7 @@
   done
 
 
-subsection {* Some random examples *}
+subsection \<open>Some random examples\<close>
 
 schematic_lemma (in LP2) "A:* c:A f:A\<rightarrow>A \<turnstile>
     \<Lambda> a:A. \<Pi> P:A\<rightarrow>*.P^c \<rightarrow> (\<Pi> x:A. P^x\<rightarrow>P^(f^x)) \<rightarrow> P^a : ?T"
@@ -211,7 +206,7 @@
 
 schematic_lemma (in LP2)
   "A:* a:A b:A \<turnstile> ?p: (\<Pi> P:A\<rightarrow>*.P^a\<rightarrow>P^b) \<rightarrow> (\<Pi> P:A\<rightarrow>*.P^b\<rightarrow>P^a)"
-  -- {* Symmetry of Leibnitz equality *}
+  -- \<open>Symmetry of Leibnitz equality\<close>
   apply (strip_asms rules)
   apply (rule lam_ss)
     apply (depth_solve1 rules)