merged
authornipkow
Sat, 22 Oct 2011 20:18:01 +0200
changeset 45247 d7f5338f0335
parent 45245 7f6c85421fa9 (diff)
parent 45246 4fbeabee6487 (current diff)
child 45248 3b7b64b194ee
merged
--- a/src/Cube/Cube.thy	Sat Oct 22 20:17:50 2011 +0200
+++ b/src/Cube/Cube.thy	Sat Oct 22 20:18:01 2011 +0200
@@ -14,69 +14,69 @@
 typedecl "context"
 typedecl typing
 
-nonterminal context' and typing'
+axiomatization
+  Abs :: "[term, term => term] => term" and
+  Prod :: "[term, term => term] => term" and
+  Trueprop :: "[context, typing] => prop" and
+  MT_context :: "context" and
+  Context :: "[typing, context] => context" and
+  star :: "term"  ("*") and
+  box :: "term"  ("\<box>") and
+  app :: "[term, term] => term"  (infixl "^" 20) and
+  Has_type :: "[term, term] => typing"
 
-consts
-  Abs :: "[term, term => term] => term"
-  Prod :: "[term, term => term] => term"
-  Trueprop :: "[context, typing] => prop"
-  MT_context :: "context"
-  Context :: "[typing, context] => context"
-  star :: "term"    ("*")
-  box :: "term"    ("[]")
-  app :: "[term, term] => term"    (infixl "^" 20)
-  Has_type :: "[term, term] => typing"
+nonterminal context' and typing'
 
 syntax
-  "\<^const>Cube.Trueprop" :: "[context', typing'] => prop"    ("(_/ |- _)")
-  "_Trueprop1" :: "typing' => prop"    ("(_)")
-  "" :: "id => context'"    ("_")
-  "" :: "var => context'"    ("_")
-  "\<^const>Cube.MT_context" :: "context'"    ("")
-  "\<^const>Cube.Context" :: "[typing', context'] => context'"    ("_ _")
-  "\<^const>Cube.Has_type" :: "[term, term] => typing'"    ("(_:/ _)" [0, 0] 5)
-  "_Lam" :: "[idt, term, term] => term"    ("(3Lam _:_./ _)" [0, 0, 0] 10)
-  "_Pi" :: "[idt, term, term] => term"    ("(3Pi _:_./ _)" [0, 0] 10)
-  "_arrow" :: "[term, term] => term"    (infixr "->" 10)
+  "_Trueprop" :: "[context', typing'] => prop"  ("(_/ \<turnstile> _)")
+  "_Trueprop1" :: "typing' => prop"  ("(_)")
+  "" :: "id => context'"  ("_")
+  "" :: "var => context'"  ("_")
+  "_MT_context" :: "context'"  ("")
+  "_Context" :: "[typing', context'] => context'"  ("_ _")
+  "_Has_type" :: "[term, term] => typing'"  ("(_:/ _)" [0, 0] 5)
+  "_Lam" :: "[idt, term, term] => term"  ("(3\<Lambda> _:_./ _)" [0, 0, 0] 10)
+  "_Pi" :: "[idt, term, term] => term"  ("(3\<Pi> _:_./ _)" [0, 0] 10)
+  "_arrow" :: "[term, term] => term"  (infixr "\<rightarrow>" 10)
 
 translations
-  ("prop") "x:X" == ("prop") "|- x:X"
-  "Lam x:A. B" == "CONST Abs(A, %x. B)"
-  "Pi x:A. B" => "CONST Prod(A, %x. B)"
-  "A -> B" => "CONST Prod(A, %_. B)"
+  "_Trueprop(G, t)" == "CONST Trueprop(G, t)"
+  ("prop") "x:X" == ("prop") "\<turnstile> x:X"
+  "_MT_context" == "CONST MT_context"
+  "_Context" == "CONST Context"
+  "_Has_type" == "CONST Has_type"
+  "\<Lambda> x:A. B" == "CONST Abs(A, %x. B)"
+  "\<Pi> x:A. B" => "CONST Prod(A, %x. B)"
+  "A \<rightarrow> B" => "CONST Prod(A, %_. B)"
 
 syntax (xsymbols)
-  "\<^const>Cube.Trueprop" :: "[context', typing'] => prop"    ("(_/ \<turnstile> _)")
-  "\<^const>Cube.box" :: "term"    ("\<box>")
-  "_Lam" :: "[idt, term, term] => term"    ("(3\<Lambda> _:_./ _)" [0, 0, 0] 10)
   "_Pi" :: "[idt, term, term] => term"    ("(3\<Pi> _:_./ _)" [0, 0] 10)
-  "_arrow" :: "[term, term] => term"    (infixr "\<rightarrow>" 10)
 
 print_translation {*
   [(@{const_syntax Prod},
     Syntax_Trans.dependent_tr' (@{syntax_const "_Pi"}, @{syntax_const "_arrow"}))]
 *}
 
-axioms
-  s_b:          "*: []"
+axiomatization where
+  s_b: "*: \<box>"  and
 
-  strip_s:      "[| A:*;  a:A ==> G |- x:X |] ==> a:A G |- x:X"
-  strip_b:      "[| A:[]; a:A ==> G |- x:X |] ==> a:A G |- x:X"
+  strip_s: "[| A:*;  a:A ==> G \<turnstile> x:X |] ==> a:A G \<turnstile> x:X" and
+  strip_b: "[| A:\<box>; a:A ==> G \<turnstile> x:X |] ==> a:A G \<turnstile> x:X" and
 
-  app:          "[| F:Prod(A, B); C:A |] ==> F^C: B(C)"
+  app: "[| F:Prod(A, B); C:A |] ==> F^C: B(C)" and
 
-  pi_ss:        "[| A:*; !!x. x:A ==> B(x):* |] ==> Prod(A, B):*"
+  pi_ss: "[| A:*; !!x. x:A ==> B(x):* |] ==> Prod(A, B):*" and
 
-  lam_ss:       "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |]
-                   ==> Abs(A, f) : Prod(A, B)"
+  lam_ss: "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |]
+            ==> Abs(A, f) : Prod(A, B)" and
 
-  beta:          "Abs(A, f)^a == f(a)"
+  beta: "Abs(A, f)^a == f(a)"
 
 lemmas simple = s_b strip_s strip_b app lam_ss pi_ss
 lemmas rules = simple
 
 lemma imp_elim:
-  assumes "f:A->B" and "a:A" and "f^a:B ==> PROP P"
+  assumes "f:A\<rightarrow>B" and "a:A" and "f^a:B ==> PROP P"
   shows "PROP P" by (rule app assms)+
 
 lemma pi_elim:
@@ -85,47 +85,68 @@
 
 
 locale L2 =
-  assumes pi_bs: "[| A:[]; !!x. x:A ==> B(x):* |] ==> Prod(A,B):*"
-    and lam_bs: "[| A:[]; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |]
+  assumes pi_bs: "[| A:\<box>; !!x. x:A ==> B(x):* |] ==> Prod(A,B):*"
+    and lam_bs: "[| A:\<box>; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |]
                    ==> Abs(A,f) : Prod(A,B)"
+begin
 
-lemmas (in L2) rules = simple lam_bs pi_bs
+lemmas rules = simple lam_bs pi_bs
+
+end
 
 
 locale Lomega =
   assumes
-    pi_bb: "[| A:[]; !!x. x:A ==> B(x):[] |] ==> Prod(A,B):[]"
-    and lam_bb: "[| A:[]; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):[] |]
+    pi_bb: "[| A:\<box>; !!x. x:A ==> B(x):\<box> |] ==> Prod(A,B):\<box>"
+    and lam_bb: "[| A:\<box>; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):\<box> |]
                    ==> Abs(A,f) : Prod(A,B)"
+begin
 
-lemmas (in Lomega) rules = simple lam_bb pi_bb
+lemmas rules = simple lam_bb pi_bb
+
+end
 
 
 locale LP =
-  assumes pi_sb: "[| A:*; !!x. x:A ==> B(x):[] |] ==> Prod(A,B):[]"
-    and lam_sb: "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):[] |]
+  assumes pi_sb: "[| A:*; !!x. x:A ==> B(x):\<box> |] ==> Prod(A,B):\<box>"
+    and lam_sb: "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):\<box> |]
                    ==> Abs(A,f) : Prod(A,B)"
+begin
 
-lemmas (in LP) rules = simple lam_sb pi_sb
+lemmas rules = simple lam_sb pi_sb
+
+end
 
 
 locale LP2 = LP + L2
+begin
 
-lemmas (in LP2) rules = simple lam_bs pi_bs lam_sb pi_sb
+lemmas rules = simple lam_bs pi_bs lam_sb pi_sb
+
+end
 
 
 locale Lomega2 = L2 + Lomega
+begin
 
-lemmas (in Lomega2) rules = simple lam_bs pi_bs lam_bb pi_bb
+lemmas rules = simple lam_bs pi_bs lam_bb pi_bb
+
+end
 
 
 locale LPomega = LP + Lomega
+begin
 
-lemmas (in LPomega) rules = simple lam_bb pi_bb lam_sb pi_sb
+lemmas rules = simple lam_bb pi_bb lam_sb pi_sb
+
+end
 
 
 locale CC = L2 + LP + Lomega
+begin
 
-lemmas (in CC) rules = simple lam_bs pi_bs lam_bb pi_bb lam_sb pi_sb
+lemmas rules = simple lam_bs pi_bs lam_bb pi_bb lam_sb pi_sb
 
 end
+
+end
--- a/src/Cube/Example.thy	Sat Oct 22 20:17:50 2011 +0200
+++ b/src/Cube/Example.thy	Sat Oct 22 20:18:01 2011 +0200
@@ -30,98 +30,98 @@
 
 subsection {* Simple types *}
 
-schematic_lemma "A:* |- A->A : ?T"
+schematic_lemma "A:* \<turnstile> A\<rightarrow>A : ?T"
   by (depth_solve rules)
 
-schematic_lemma "A:* |- Lam a:A. a : ?T"
+schematic_lemma "A:* \<turnstile> \<Lambda> a:A. a : ?T"
   by (depth_solve rules)
 
-schematic_lemma "A:* B:* b:B |- Lam x:A. b : ?T"
+schematic_lemma "A:* B:* b:B \<turnstile> \<Lambda> x:A. b : ?T"
   by (depth_solve rules)
 
-schematic_lemma "A:* b:A |- (Lam a:A. a)^b: ?T"
+schematic_lemma "A:* b:A \<turnstile> (\<Lambda> a:A. a)^b: ?T"
   by (depth_solve rules)
 
-schematic_lemma "A:* B:* c:A b:B |- (Lam x:A. b)^ c: ?T"
+schematic_lemma "A:* B:* c:A b:B \<turnstile> (\<Lambda> x:A. b)^ c: ?T"
   by (depth_solve rules)
 
-schematic_lemma "A:* B:* |- Lam a:A. Lam b:B. a : ?T"
+schematic_lemma "A:* B:* \<turnstile> \<Lambda> a:A. \<Lambda> b:B. a : ?T"
   by (depth_solve rules)
 
 
 subsection {* Second-order types *}
 
-schematic_lemma (in L2) "|- Lam A:*. Lam a:A. a : ?T"
+schematic_lemma (in L2) "\<turnstile> \<Lambda> A:*. \<Lambda> a:A. a : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in L2) "A:* |- (Lam B:*.Lam b:B. b)^A : ?T"
+schematic_lemma (in L2) "A:* \<turnstile> (\<Lambda> B:*.\<Lambda> b:B. b)^A : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in L2) "A:* b:A |- (Lam B:*.Lam b:B. b) ^ A ^ b: ?T"
+schematic_lemma (in L2) "A:* b:A \<turnstile> (\<Lambda> B:*.\<Lambda> b:B. b) ^ A ^ b: ?T"
   by (depth_solve rules)
 
-schematic_lemma (in L2) "|- Lam B:*.Lam a:(Pi A:*.A).a ^ ((Pi A:*.A)->B) ^ a: ?T"
+schematic_lemma (in L2) "\<turnstile> \<Lambda> B:*.\<Lambda> a:(\<Pi> A:*.A).a ^ ((\<Pi> A:*.A)\<rightarrow>B) ^ a: ?T"
   by (depth_solve rules)
 
 
 subsection {* Weakly higher-order propositional logic *}
 
-schematic_lemma (in Lomega) "|- Lam A:*.A->A : ?T"
+schematic_lemma (in Lomega) "\<turnstile> \<Lambda> A:*.A\<rightarrow>A : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in Lomega) "B:* |- (Lam A:*.A->A) ^ B : ?T"
+schematic_lemma (in Lomega) "B:* \<turnstile> (\<Lambda> A:*.A\<rightarrow>A) ^ B : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in Lomega) "B:* b:B |- (Lam y:B. b): ?T"
+schematic_lemma (in Lomega) "B:* b:B \<turnstile> (\<Lambda> y:B. b): ?T"
   by (depth_solve rules)
 
-schematic_lemma (in Lomega) "A:* F:*->* |- F^(F^A): ?T"
+schematic_lemma (in Lomega) "A:* F:*\<rightarrow>* \<turnstile> F^(F^A): ?T"
   by (depth_solve rules)
 
-schematic_lemma (in Lomega) "A:* |- Lam F:*->*.F^(F^A): ?T"
+schematic_lemma (in Lomega) "A:* \<turnstile> \<Lambda> F:*\<rightarrow>*.F^(F^A): ?T"
   by (depth_solve rules)
 
 
 subsection {* LP *}
 
-schematic_lemma (in LP) "A:* |- A -> * : ?T"
+schematic_lemma (in LP) "A:* \<turnstile> A \<rightarrow> * : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in LP) "A:* P:A->* a:A |- P^a: ?T"
+schematic_lemma (in LP) "A:* P:A\<rightarrow>* a:A \<turnstile> P^a: ?T"
   by (depth_solve rules)
 
-schematic_lemma (in LP) "A:* P:A->A->* a:A |- Pi a:A. P^a^a: ?T"
+schematic_lemma (in LP) "A:* P:A\<rightarrow>A\<rightarrow>* a:A \<turnstile> \<Pi> a:A. P^a^a: ?T"
   by (depth_solve rules)
 
-schematic_lemma (in LP) "A:* P:A->* Q:A->* |- Pi a:A. P^a -> Q^a: ?T"
+schematic_lemma (in LP) "A:* P:A\<rightarrow>* Q:A\<rightarrow>* \<turnstile> \<Pi> a:A. P^a \<rightarrow> Q^a: ?T"
   by (depth_solve rules)
 
-schematic_lemma (in LP) "A:* P:A->* |- Pi a:A. P^a -> P^a: ?T"
+schematic_lemma (in LP) "A:* P:A\<rightarrow>* \<turnstile> \<Pi> a:A. P^a \<rightarrow> P^a: ?T"
   by (depth_solve rules)
 
-schematic_lemma (in LP) "A:* P:A->* |- Lam a:A. Lam x:P^a. x: ?T"
+schematic_lemma (in LP) "A:* P:A\<rightarrow>* \<turnstile> \<Lambda> a:A. \<Lambda> x:P^a. x: ?T"
   by (depth_solve rules)
 
-schematic_lemma (in LP) "A:* P:A->* Q:* |- (Pi a:A. P^a->Q) -> (Pi a:A. P^a) -> Q : ?T"
+schematic_lemma (in LP) "A:* P:A\<rightarrow>* Q:* \<turnstile> (\<Pi> a:A. P^a\<rightarrow>Q) \<rightarrow> (\<Pi> a:A. P^a) \<rightarrow> Q : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in LP) "A:* P:A->* Q:* a0:A |-
-        Lam x:Pi a:A. P^a->Q. Lam y:Pi a:A. P^a. x^a0^(y^a0): ?T"
+schematic_lemma (in LP) "A:* P:A\<rightarrow>* Q:* a0:A \<turnstile>
+        \<Lambda> x:\<Pi> a:A. P^a\<rightarrow>Q. \<Lambda> y:\<Pi> a:A. P^a. x^a0^(y^a0): ?T"
   by (depth_solve rules)
 
 
 subsection {* Omega-order types *}
 
-schematic_lemma (in L2) "A:* B:* |- Pi C:*.(A->B->C)->C : ?T"
+schematic_lemma (in L2) "A:* B:* \<turnstile> \<Pi> C:*.(A\<rightarrow>B\<rightarrow>C)\<rightarrow>C : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in Lomega2) "|- Lam A:*.Lam B:*.Pi C:*.(A->B->C)->C : ?T"
+schematic_lemma (in Lomega2) "\<turnstile> \<Lambda> A:*.\<Lambda> B:*.\<Pi> C:*.(A\<rightarrow>B\<rightarrow>C)\<rightarrow>C : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in Lomega2) "|- Lam A:*.Lam B:*.Lam x:A. Lam y:B. x : ?T"
+schematic_lemma (in Lomega2) "\<turnstile> \<Lambda> A:*.\<Lambda> B:*.\<Lambda> x:A. \<Lambda> y:B. x : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in Lomega2) "A:* B:* |- ?p : (A->B) -> ((B->Pi P:*.P)->(A->Pi P:*.P))"
+schematic_lemma (in Lomega2) "A:* B:* \<turnstile> ?p : (A\<rightarrow>B) \<rightarrow> ((B\<rightarrow>\<Pi> P:*.P)\<rightarrow>(A\<rightarrow>\<Pi> P:*.P))"
   apply (strip_asms rules)
   apply (rule lam_ss)
     apply (depth_solve1 rules)
@@ -145,15 +145,15 @@
 
 subsection {* Second-order Predicate Logic *}
 
-schematic_lemma (in LP2) "A:* P:A->* |- Lam a:A. P^a->(Pi A:*.A) : ?T"
+schematic_lemma (in LP2) "A:* P:A\<rightarrow>* \<turnstile> \<Lambda> a:A. P^a\<rightarrow>(\<Pi> A:*.A) : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in LP2) "A:* P:A->A->* |-
-    (Pi a:A. Pi b:A. P^a^b->P^b^a->Pi P:*.P) -> Pi a:A. P^a^a->Pi P:*.P : ?T"
+schematic_lemma (in LP2) "A:* P:A\<rightarrow>A\<rightarrow>* \<turnstile>
+    (\<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 : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in LP2) "A:* P:A->A->* |-
-    ?p: (Pi a:A. Pi b:A. P^a^b->P^b^a->Pi P:*.P) -> Pi a:A. P^a^a->Pi P:*.P"
+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: *}
   apply (strip_asms rules)
   apply (rule lam_ss)
@@ -174,22 +174,22 @@
 
 subsection {* LPomega *}
 
-schematic_lemma (in LPomega) "A:* |- Lam P:A->A->*.Lam a:A. P^a^a : ?T"
+schematic_lemma (in LPomega) "A:* \<turnstile> \<Lambda> P:A\<rightarrow>A\<rightarrow>*.\<Lambda> a:A. P^a^a : ?T"
   by (depth_solve rules)
 
-schematic_lemma (in LPomega) "|- Lam A:*.Lam P:A->A->*.Lam a:A. P^a^a : ?T"
+schematic_lemma (in LPomega) "\<turnstile> \<Lambda> A:*.\<Lambda> P:A\<rightarrow>A\<rightarrow>*.\<Lambda> a:A. P^a^a : ?T"
   by (depth_solve rules)
 
 
 subsection {* Constructions *}
 
-schematic_lemma (in CC) "|- Lam A:*.Lam P:A->*.Lam a:A. P^a->Pi P:*.P: ?T"
+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)
 
-schematic_lemma (in CC) "|- Lam A:*.Lam P:A->*.Pi a:A. P^a: ?T"
+schematic_lemma (in CC) "\<turnstile> \<Lambda> A:*.\<Lambda> P:A\<rightarrow>*.\<Pi> a:A. P^a: ?T"
   by (depth_solve rules)
 
-schematic_lemma (in CC) "A:* P:A->* a:A |- ?p : (Pi a:A. P^a)->P^a"
+schematic_lemma (in CC) "A:* P:A\<rightarrow>* a:A \<turnstile> ?p : (\<Pi> a:A. P^a)\<rightarrow>P^a"
   apply (strip_asms rules)
   apply (rule lam_ss)
     apply (depth_solve1 rules)
@@ -201,23 +201,23 @@
 
 subsection {* Some random examples *}
 
-schematic_lemma (in LP2) "A:* c:A f:A->A |-
-    Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T"
+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"
   by (depth_solve rules)
 
-schematic_lemma (in CC) "Lam A:*.Lam c:A. Lam f:A->A.
-    Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T"
+schematic_lemma (in CC) "\<Lambda> A:*.\<Lambda> c:A. \<Lambda> f:A\<rightarrow>A.
+    \<Lambda> a:A. \<Pi> P:A\<rightarrow>*.P^c \<rightarrow> (\<Pi> x:A. P^x\<rightarrow>P^(f^x)) \<rightarrow> P^a : ?T"
   by (depth_solve rules)
 
 schematic_lemma (in LP2)
-  "A:* a:A b:A |- ?p: (Pi P:A->*.P^a->P^b) -> (Pi P:A->*.P^b->P^a)"
+  "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 *}
   apply (strip_asms rules)
   apply (rule lam_ss)
     apply (depth_solve1 rules)
    prefer 2
    apply (depth_solve1 rules)
-  apply (erule_tac a = "Lam x:A. Pi Q:A->*.Q^x->Q^a" in pi_elim)
+  apply (erule_tac a = "\<Lambda> x:A. \<Pi> Q:A\<rightarrow>*.Q^x\<rightarrow>Q^a" in pi_elim)
    apply (depth_solve1 rules)
   apply (unfold beta)
   apply (erule imp_elim)
--- a/src/Cube/ROOT.ML	Sat Oct 22 20:17:50 2011 +0200
+++ b/src/Cube/ROOT.ML	Sat Oct 22 20:18:01 2011 +0200
@@ -5,4 +5,4 @@
 The Lambda-Cube a la Barendregt.
 *)
 
-use_thys ["Cube", "Example"];
+use_thys ["Example"];
--- a/src/Pure/Concurrent/counter.scala	Sat Oct 22 20:17:50 2011 +0200
+++ b/src/Pure/Concurrent/counter.scala	Sat Oct 22 20:18:01 2011 +0200
@@ -12,9 +12,10 @@
 object Counter
 {
   type ID = Long
+  def apply(): Counter = new Counter
 }
 
-class Counter
+class Counter private
 {
   private var count: Counter.ID = 0
 
--- a/src/Pure/General/path.scala	Sat Oct 22 20:17:50 2011 +0200
+++ b/src/Pure/General/path.scala	Sat Oct 22 20:18:01 2011 +0200
@@ -15,7 +15,7 @@
 {
   /* path elements */
 
-  private sealed abstract class Elem
+  sealed abstract class Elem
   private case class Root(val name: String) extends Elem
   private case class Basic(val name: String) extends Elem
   private case class Variable(val name: String) extends Elem
@@ -60,15 +60,12 @@
 
   /* path constructors */
 
-  private def apply(xs: List[Elem]): Path =
-    new Path { override val elems = xs }
-
-  val current: Path = Path(Nil)
-  val root: Path = Path(List(Root("")))
-  def named_root(s: String): Path = Path(List(root_elem(s)))
-  def basic(s: String): Path = Path(List(basic_elem(s)))
-  def variable(s: String): Path = Path(List(variable_elem(s)))
-  val parent: Path = Path(List(Parent))
+  val current: Path = new Path(Nil)
+  val root: Path = new Path(List(Root("")))
+  def named_root(s: String): Path = new Path(List(root_elem(s)))
+  def basic(s: String): Path = new Path(List(basic_elem(s)))
+  def variable(s: String): Path = new Path(List(variable_elem(s)))
+  val parent: Path = new Path(List(Parent))
 
 
   /* explode */
@@ -93,7 +90,7 @@
       else if (r == 1) (List(Root("")), es)
       else if (es.isEmpty) (List(Root("")), Nil)
       else (List(root_elem(es.head)), es.tail)
-    Path(norm_elems(explode_elems(raw_elems) ++ roots))
+    new Path(norm_elems(explode_elems(raw_elems) ++ roots))
   }
 
   def split(str: String): List[Path] =
@@ -101,15 +98,13 @@
 }
 
 
-class Path
+class Path private(private val elems: List[Path.Elem]) // reversed elements
 {
-  protected val elems: List[Path.Elem] = Nil   // reversed elements
-
   def is_current: Boolean = elems.isEmpty
   def is_absolute: Boolean = !elems.isEmpty && elems.last.isInstanceOf[Path.Root]
   def is_basic: Boolean = elems match { case List(Path.Basic(_)) => true case _ => false }
 
-  def +(other: Path): Path = Path((other.elems :\ elems)(Path.apply_elem))
+  def +(other: Path): Path = new Path((other.elems :\ elems)(Path.apply_elem))
 
 
   /* implode */
@@ -128,12 +123,12 @@
 
   private def split_path: (Path, String) =
     elems match {
-      case Path.Basic(s) :: xs => (Path(xs), s)
+      case Path.Basic(s) :: xs => (new Path(xs), s)
       case _ => error("Cannot split path into dir/base: " + toString)
     }
 
   def dir: Path = split_path._1
-  def base: Path = Path(List(Path.Basic(split_path._2)))
+  def base: Path = new Path(List(Path.Basic(split_path._2)))
 
   def ext(e: String): Path =
     if (e == "") this
@@ -165,6 +160,6 @@
         case x => List(x)
       }
 
-    Path(Path.norm_elems(elems.map(eval).flatten))
+    new Path(Path.norm_elems(elems.map(eval).flatten))
   }
 }
--- a/src/Pure/PIDE/document.scala	Sat Oct 22 20:17:50 2011 +0200
+++ b/src/Pure/PIDE/document.scala	Sat Oct 22 20:18:01 2011 +0200
@@ -23,7 +23,7 @@
   type Exec_ID = ID
 
   val no_id: ID = 0
-  val new_id = new Counter
+  val new_id = Counter()
 
 
 
--- a/src/Pure/PIDE/text.scala	Sat Oct 22 20:17:50 2011 +0200
+++ b/src/Pure/PIDE/text.scala	Sat Oct 22 20:18:01 2011 +0200
@@ -51,12 +51,19 @@
     def overlaps(that: Range): Boolean = this.contains(that.start) || that.contains(this.start)
     def compare(that: Range): Int = if (overlaps(that)) 0 else this.start compare that.start
 
+    def apart(that: Range): Boolean =
+      (this.start max that.start) > (this.stop min that.stop)
+
     def restrict(that: Range): Range =
       Range(this.start max that.start, this.stop min that.stop)
 
     def try_restrict(that: Range): Option[Range] =
-      try { Some (restrict(that)) }
-      catch { case ERROR(_) => None }
+      if (this apart that) None
+      else Some(restrict(that))
+
+    def try_join(that: Range): Option[Range] =
+      if (this apart that) None
+      else Some(Range(this.start min that.start, this.stop max that.stop))
   }
 
 
@@ -68,33 +75,33 @@
 
     def apply(ranges: Seq[Range]): Perspective =
     {
-      val sorted_ranges = ranges.toArray
-      Sorting.quickSort(sorted_ranges)(Range.Ordering)
-
       val result = new mutable.ListBuffer[Text.Range]
       var last: Option[Text.Range] = None
-      for (range <- sorted_ranges)
+      def ship(next: Option[Range]) { result ++= last; last = next }
+
+      for (range <- ranges.sortBy(_.start))
       {
         last match {
-          case Some(last_range)
-          if ((last_range overlaps range) || last_range.stop == range.start) =>
-            last = Some(Text.Range(last_range.start, range.stop))
-          case _ =>
-            result ++= last
-            last = Some(range)
+          case None => ship(Some(range))
+          case Some(last_range) =>
+            last_range.try_join(range) match {
+              case None => ship(Some(range))
+              case joined => last = joined
+            }
         }
       }
-      result ++= last
+      ship(None)
       new Perspective(result.toList)
     }
   }
 
-  sealed case class Perspective(ranges: List[Range]) // visible text partitioning in canonical order
+  class Perspective private(val ranges: List[Range]) // visible text partitioning in canonical order
   {
     def is_empty: Boolean = ranges.isEmpty
     def range: Range =
       if (is_empty) Range(0)
       else Range(ranges.head.start, ranges.last.stop)
+    override def toString = ranges.toString
   }
 
 
--- a/src/Pure/System/system_channel.scala	Sat Oct 22 20:17:50 2011 +0200
+++ b/src/Pure/System/system_channel.scala	Sat Oct 22 20:18:01 2011 +0200
@@ -30,7 +30,7 @@
 
 object Fifo_Channel
 {
-  private val next_fifo = new Counter
+  private val next_fifo = Counter()
 }
 
 class Fifo_Channel extends System_Channel
--- a/src/Tools/jEdit/README_BUILD	Sat Oct 22 20:17:50 2011 +0200
+++ b/src/Tools/jEdit/README_BUILD	Sat Oct 22 20:18:01 2011 +0200
@@ -4,7 +4,7 @@
 * Official Java JDK from Sun/Oracle/Apple, e.g. 1.6.0_27
   http://www.oracle.com/technetwork/java/javase/downloads/index.html
 
-* Scala 2.8.2.final
+* Scala 2.8.2.final (experimental support for 2.9.1.final)
   http://www.scala-lang.org
 
 * Auxiliary jedit_build component