# HG changeset patch # User wenzelm # Date 1412710104 -7200 # Node ID 4f169d2cf6f3b30be52b1888389d87b1139aebe4 # Parent 4257a7f2bf397cc5d48969e951cc9c0ee71daa73 more cartouches; diff -r 4257a7f2bf39 -r 4f169d2cf6f3 src/Cube/Cube.thy --- a/src/Cube/Cube.thy Tue Oct 07 21:11:18 2014 +0200 +++ b/src/Cube/Cube.thy Tue Oct 07 21:28:24 2014 +0200 @@ -2,7 +2,7 @@ Author: Tobias Nipkow *) -header {* Barendregt's Lambda-Cube *} +header \Barendregt's Lambda-Cube\ theory Cube imports Pure @@ -54,10 +54,10 @@ syntax (xsymbols) "_Pi" :: "[idt, term, term] => term" ("(3\ _:_./ _)" [0, 0] 10) -print_translation {* +print_translation \ [(@{const_syntax Prod}, fn _ => Syntax_Trans.dependent_tr' (@{syntax_const "_Pi"}, @{syntax_const "_arrow"}))] -*} +\ axiomatization where s_b: "*: \" and diff -r 4257a7f2bf39 -r 4f169d2cf6f3 src/Cube/Example.thy --- 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 \Lambda Cube Examples\ theory Example imports Cube begin -text {* - Examples taken from: +text \Examples taken from: H. Barendregt. Introduction to Generalised Type Systems. - J. Functional Programming. -*} + J. Functional Programming.\ -method_setup depth_solve = {* - Attrib.thms >> (fn thms => K (METHOD (fn facts => - (DEPTH_SOLVE (HEADGOAL (ares_tac (facts @ thms))))))) -*} +method_setup depth_solve = + \Attrib.thms >> (fn thms => K (METHOD (fn facts => + (DEPTH_SOLVE (HEADGOAL (ares_tac (facts @ thms)))))))\ -method_setup depth_solve1 = {* - Attrib.thms >> (fn thms => K (METHOD (fn facts => - (DEPTH_SOLVE_1 (HEADGOAL (ares_tac (facts @ thms))))))) -*} +method_setup depth_solve1 = + \Attrib.thms >> (fn thms => K (METHOD (fn facts => + (DEPTH_SOLVE_1 (HEADGOAL (ares_tac (facts @ thms)))))))\ -method_setup strip_asms = {* - Attrib.thms >> (fn thms => K (METHOD (fn facts => +method_setup strip_asms = + \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)))))\ -subsection {* Simple types *} +subsection \Simple types\ schematic_lemma "A:* \ A\A : ?T" by (depth_solve rules) @@ -49,7 +44,7 @@ by (depth_solve rules) -subsection {* Second-order types *} +subsection \Second-order types\ schematic_lemma (in L2) "\ \ A:*. \ a:A. a : ?T" by (depth_solve rules) @@ -64,7 +59,7 @@ by (depth_solve rules) -subsection {* Weakly higher-order propositional logic *} +subsection \Weakly higher-order propositional logic\ schematic_lemma (in Lomega) "\ \ A:*.A\A : ?T" by (depth_solve rules) @@ -82,7 +77,7 @@ by (depth_solve rules) -subsection {* LP *} +subsection \LP\ schematic_lemma (in LP) "A:* \ A \ * : ?T" by (depth_solve rules) @@ -110,7 +105,7 @@ by (depth_solve rules) -subsection {* Omega-order types *} +subsection \Omega-order types\ schematic_lemma (in L2) "A:* B:* \ \ C:*.(A\B\C)\C : ?T" by (depth_solve rules) @@ -143,7 +138,7 @@ done -subsection {* Second-order Predicate Logic *} +subsection \Second-order Predicate Logic\ schematic_lemma (in LP2) "A:* P:A\* \ \ a:A. P^a\(\ A:*.A) : ?T" by (depth_solve rules) @@ -154,7 +149,7 @@ schematic_lemma (in LP2) "A:* P:A\A\* \ ?p: (\ a:A. \ b:A. P^a^b\P^b^a\\ P:*.P) \ \ a:A. P^a^a\\ P:*.P" - -- {* Antisymmetry implies irreflexivity: *} + -- \Antisymmetry implies irreflexivity:\ apply (strip_asms rules) apply (rule lam_ss) apply (depth_solve1 rules) @@ -172,7 +167,7 @@ done -subsection {* LPomega *} +subsection \LPomega\ schematic_lemma (in LPomega) "A:* \ \ P:A\A\*.\ a:A. P^a^a : ?T" by (depth_solve rules) @@ -181,7 +176,7 @@ by (depth_solve rules) -subsection {* Constructions *} +subsection \Constructions\ schematic_lemma (in CC) "\ \ A:*.\ P:A\*.\ a:A. P^a\\ P:*.P: ?T" by (depth_solve rules) @@ -199,7 +194,7 @@ done -subsection {* Some random examples *} +subsection \Some random examples\ schematic_lemma (in LP2) "A:* c:A f:A\A \ \ a:A. \ P:A\*.P^c \ (\ x:A. P^x\P^(f^x)) \ P^a : ?T" @@ -211,7 +206,7 @@ schematic_lemma (in LP2) "A:* a:A b:A \ ?p: (\ P:A\*.P^a\P^b) \ (\ P:A\*.P^b\P^a)" - -- {* Symmetry of Leibnitz equality *} + -- \Symmetry of Leibnitz equality\ apply (strip_asms rules) apply (rule lam_ss) apply (depth_solve1 rules)