clarified signature;
authorwenzelm
Sat, 09 Apr 2022 15:28:55 +0200
changeset 75436 40630fec3b5d
parent 75435 c8087e6bd2ce
child 75437 7b417c578b19
clarified signature;
src/Pure/PIDE/xml.scala
src/Pure/Thy/export_theory.scala
src/Pure/term_xml.scala
--- a/src/Pure/PIDE/xml.scala	Sat Apr 09 14:51:54 2022 +0200
+++ b/src/Pure/PIDE/xml.scala	Sat Apr 09 15:28:55 2022 +0200
@@ -333,7 +333,7 @@
 
   object Decode {
     type T[A] = XML.Body => A
-    type V[A] = (List[String], XML.Body) => A
+    type V[A] = PartialFunction[(List[String], XML.Body), A]
     type P[A] = PartialFunction[List[String], A]
 
 
--- a/src/Pure/Thy/export_theory.scala	Sat Apr 09 14:51:54 2022 +0200
+++ b/src/Pure/Thy/export_theory.scala	Sat Apr 09 15:28:55 2022 +0200
@@ -264,13 +264,12 @@
 
   def decode_syntax: XML.Decode.T[Syntax] =
     XML.Decode.variant(List(
-      { case (Nil, Nil) => No_Syntax case _ => ??? },
-      { case (List(delim), Nil) => Prefix(delim) case _ => ??? },
+      { case (Nil, Nil) => No_Syntax },
+      { case (List(delim), Nil) => Prefix(delim) },
       { case (Nil, body) =>
           import XML.Decode._
           val (ass, delim, pri) = triple(int, string, int)(body)
-          Infix(Assoc(ass), delim, pri)
-        case _ => ??? }))
+          Infix(Assoc(ass), delim, pri) }))
 
 
   /* types */
@@ -685,11 +684,11 @@
   val decode_recursion: XML.Decode.T[Recursion] = {
     import XML.Decode._
     variant(List(
-      { case (Nil, a) => Primrec(list(string)(a)) case _ => ??? },
-      { case (Nil, Nil) => Recdef case _ => ??? },
-      { case (Nil, a) => Primcorec(list(string)(a)) case _ => ??? },
-      { case (Nil, Nil) => Corec case _ => ??? },
-      { case (Nil, Nil) => Unknown_Recursion case _ => ??? }))
+      { case (Nil, a) => Primrec(list(string)(a)) },
+      { case (Nil, Nil) => Recdef },
+      { case (Nil, a) => Primcorec(list(string)(a)) },
+      { case (Nil, Nil) => Corec },
+      { case (Nil, Nil) => Unknown_Recursion }))
   }
 
 
@@ -714,10 +713,10 @@
   val decode_rough_classification: XML.Decode.T[Rough_Classification] = {
     import XML.Decode._
     variant(List(
-      { case (Nil, a) => Equational(decode_recursion(a)) case _ => ??? },
-      { case (Nil, Nil) => Inductive case _ => ??? },
-      { case (Nil, Nil) => Co_Inductive case _ => ??? },
-      { case (Nil, Nil) => Unknown case _ => ??? }))
+      { case (Nil, a) => Equational(decode_recursion(a)) },
+      { case (Nil, Nil) => Inductive },
+      { case (Nil, Nil) => Co_Inductive },
+      { case (Nil, Nil) => Unknown }))
   }
 
 
--- a/src/Pure/term_xml.scala	Sat Apr 09 14:51:54 2022 +0200
+++ b/src/Pure/term_xml.scala	Sat Apr 09 15:28:55 2022 +0200
@@ -48,20 +48,20 @@
 
     def typ: T[Typ] =
       variant[Typ](List(
-        { case (List(a), b) => Type(a, list(typ)(b)) case _ => ??? },
-        { case (List(a), b) => TFree(a, sort(b)) case _ => ??? },
+        { case (List(a), b) => Type(a, list(typ)(b)) },
+        { case (List(a), b) => TFree(a, sort(b)) },
         { case (a, b) => TVar(indexname(a), sort(b)) }))
 
     val typ_body: T[Typ] = { case Nil => dummyT case body => typ(body) }
 
     def term: T[Term] =
       variant[Term](List(
-        { case (List(a), b) => Const(a, list(typ)(b)) case _ => ??? },
-        { case (List(a), b) => Free(a, typ_body(b)) case _ => ??? },
+        { case (List(a), b) => Const(a, list(typ)(b)) },
+        { case (List(a), b) => Free(a, typ_body(b)) },
         { case (a, b) => Var(indexname(a), typ_body(b)) },
-        { case (Nil, a) => Bound(int(a)) case _ => ??? },
-        { case (List(a), b) => val (c, d) = pair(typ, term)(b); Abs(a, c, d) case _ => ??? },
-        { case (Nil, a) => val (b, c) = pair(term, term)(a); App(b, c) case _ => ??? }))
+        { case (Nil, a) => Bound(int(a)) },
+        { case (List(a), b) => val (c, d) = pair(typ, term)(b); Abs(a, c, d) },
+        { case (Nil, a) => val (b, c) = pair(term, term)(a); App(b, c) }))
 
     def term_env(env: Map[String, Typ]): T[Term] = {
       def env_type(x: String, t: Typ): Typ =
@@ -69,12 +69,12 @@
 
       def term: T[Term] =
         variant[Term](List(
-          { case (List(a), b) => Const(a, list(typ)(b)) case _ => ??? },
-          { case (List(a), b) => Free(a, env_type(a, typ_body(b))) case _ => ??? },
+          { case (List(a), b) => Const(a, list(typ)(b)) },
+          { case (List(a), b) => Free(a, env_type(a, typ_body(b))) },
           { case (a, b) => Var(indexname(a), typ_body(b)) },
-          { case (Nil, a) => Bound(int(a)) case _ => ??? },
-          { case (List(a), b) => val (c, d) = pair(typ, term)(b); Abs(a, c, d) case _ => ??? },
-          { case (Nil, a) => val (b, c) = pair(term, term)(a); App(b, c) case _ => ??? }))
+          { case (Nil, a) => Bound(int(a)) },
+          { case (List(a), b) => val (c, d) = pair(typ, term)(b); Abs(a, c, d) },
+          { case (Nil, a) => val (b, c) = pair(term, term)(a); App(b, c) }))
       term
     }
 
@@ -82,17 +82,17 @@
       val term = term_env(env)
       def proof: T[Proof] =
         variant[Proof](List(
-          { case (Nil, Nil) => MinProof case _ => ??? },
-          { case (Nil, a) => PBound(int(a)) case _ => ??? },
-          { case (List(a), b) => val (c, d) = pair(typ, proof)(b); Abst(a, c, d) case _ => ??? },
-          { case (List(a), b) => val (c, d) = pair(term, proof)(b); AbsP(a, c, d) case _ => ??? },
-          { case (Nil, a) => val (b, c) = pair(proof, term)(a); Appt(b, c) case _ => ??? },
-          { case (Nil, a) => val (b, c) = pair(proof, proof)(a); AppP(b, c) case _ => ??? },
-          { case (Nil, a) => Hyp(term(a)) case _ => ??? },
-          { case (List(a), b) => PAxm(a, list(typ)(b)) case _ => ??? },
-          { case (List(a), b) => PClass(typ(b), a) case _ => ??? },
-          { case (List(a), b) => val (c, d) = pair(term, list(typ))(b); Oracle(a, c, d) case _ => ??? },
-          { case (List(a, b, c), d) => PThm(long_atom(a), b, c, list(typ)(d)) case _ => ??? }))
+          { case (Nil, Nil) => MinProof },
+          { case (Nil, a) => PBound(int(a)) },
+          { case (List(a), b) => val (c, d) = pair(typ, proof)(b); Abst(a, c, d) },
+          { case (List(a), b) => val (c, d) = pair(term, proof)(b); AbsP(a, c, d) },
+          { case (Nil, a) => val (b, c) = pair(proof, term)(a); Appt(b, c) },
+          { case (Nil, a) => val (b, c) = pair(proof, proof)(a); AppP(b, c) },
+          { case (Nil, a) => Hyp(term(a)) },
+          { case (List(a), b) => PAxm(a, list(typ)(b)) },
+          { case (List(a), b) => PClass(typ(b), a) },
+          { case (List(a), b) => val (c, d) = pair(term, list(typ))(b); Oracle(a, c, d) },
+          { case (List(a, b, c), d) => PThm(long_atom(a), b, c, list(typ)(d)) }))
       proof
     }