simultaneous use_thys;
authorwenzelm
Tue, 31 Jul 2007 23:23:28 +0200
changeset 24106 f2965bf954dc
parent 24105 af364e2b4048
child 24107 fecafd71e758
simultaneous use_thys;
src/CCL/ROOT.ML
src/CCL/ex/ROOT.ML
src/CTT/ex/ROOT.ML
src/Cube/ROOT.ML
src/HOL/Auth/ROOT.ML
src/HOLCF/FOCUS/ROOT.ML
src/HOLCF/IMP/ROOT.ML
src/HOLCF/ex/ROOT.ML
src/LCF/ex/ROOT.ML
src/Sequents/LK/ROOT.ML
src/Sequents/ROOT.ML
--- a/src/CCL/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/CCL/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -6,13 +6,10 @@
 Classical Computational Logic based on First-Order Logic.
 *)
 
-val banner = "Classical Computational Logic (in FOL)";
-writeln banner;
-
 set eta_contract;
 
 (* CCL - a computational logic for an untyped functional language *)
 (*                       with evaluation to weak head-normal form *)
 
-use_thy "Wfd";
-use_thy "Fix";
+use_thys ["Wfd", "Fix"];
+
--- a/src/CCL/ex/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/CCL/ex/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -6,7 +6,4 @@
 Examples for Classical Computational Logic.
 *)
 
-time_use_thy "Nat";
-time_use_thy "List";
-time_use_thy "Stream";
-time_use_thy "Flag";
+use_thys ["Nat", "List", "Stream", "Flag"];
--- a/src/CTT/ex/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/CTT/ex/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -6,7 +6,4 @@
 Examples for Constructive Type Theory. 
 *)
 
-use_thy "Typechecking";
-use_thy "Elimination";
-use_thy "Equality";
-use_thy "Synthesis";
+use_thys ["Typechecking", "Elimination", "Equality", "Synthesis"];
--- a/src/Cube/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/Cube/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -6,8 +6,4 @@
 The Lambda-Cube a la Barendregt.
 *)
 
-val banner = "Barendregt's Lambda-Cube";
-writeln banner;
-
-use_thy "Cube";
-use_thy "Example";
+use_thys ["Cube", "Example"];
--- a/src/HOL/Auth/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/HOL/Auth/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -1,4 +1,4 @@
-(*  Title:      HOL/Auth/ROOT
+(*  Title:      HOL/Auth/ROOT.ML
     ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1996  University of Cambridge
@@ -8,41 +8,46 @@
 
 no_document use_thy "NatPair";
 
-set timing;
+use_thys [
 
 (* Conventional protocols: rely on 
    conventional Message, Event and Public *)
 
 (*Shared-key protocols*)
-time_use_thy "NS_Shared";
-time_use_thy "Kerberos_BAN";
-time_use_thy "Kerberos_BAN_Gets";
-time_use_thy "KerberosIV";
-time_use_thy "KerberosIV_Gets";
-time_use_thy "KerberosV";
-time_use_thy "OtwayRees";
-time_use_thy "OtwayRees_AN";
-time_use_thy "OtwayRees_Bad";
-time_use_thy "OtwayReesBella";
-time_use_thy "WooLam";
-time_use_thy "Recur";
-time_use_thy "Yahalom";
-time_use_thy "Yahalom2";
-time_use_thy "Yahalom_Bad";
-time_use_thy "ZhouGollmann";
+  "NS_Shared",
+  "Kerberos_BAN",
+  "Kerberos_BAN_Gets",
+  "KerberosIV",
+  "KerberosIV_Gets",
+  "KerberosV",
+  "OtwayRees",
+  "OtwayRees_AN",
+  "OtwayRees_Bad",
+  "OtwayReesBella",
+  "WooLam",
+  "Recur",
+  "Yahalom",
+  "Yahalom2",
+  "Yahalom_Bad",
+  "ZhouGollmann",
 
 (*Public-key protocols*)
-time_use_thy "NS_Public_Bad";
-time_use_thy "NS_Public";
-time_use_thy "TLS";
-time_use_thy "CertifiedEmail";
+  "NS_Public_Bad",
+  "NS_Public",
+  "TLS",
+  "CertifiedEmail",
 
 (*Smartcard protocols: rely on conventional Message and on new
   EventSC and Smartcard *)
 
-with_path "Smartcard" time_use_thy "ShoupRubin";
-with_path "Smartcard" time_use_thy "ShoupRubinBella";
+  "Smartcard/ShoupRubin",
+  "Smartcard/ShoupRubinBella",
 
 (*Blanqui's "guard" concept: protocol-independent secrecy*)
-with_path "Guard" (app time_use_thy)
-  ["P1", "P2", "Guard_NS_Public", "Guard_OtwayRees", "Guard_Yahalom", "Proto"];
+  "Guard/P1",
+  "Guard/P2",
+  "Guard/Guard_NS_Public",
+  "Guard/Guard_OtwayRees",
+  "Guard/Guard_Yahalom",
+  "Guard/Proto"
+];
--- a/src/HOLCF/FOCUS/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/HOLCF/FOCUS/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -6,9 +6,4 @@
 See README.html for further information.
 *)
 
-val banner = "HOLCF/FOCUS";
-writeln banner;
-
-with_path "~~/src/HOLCF/ex" use_thy "Fstreams";
-use_thy "FOCUS";
-use_thy "Buffer_adm";
+use_thys ["Fstreams", "FOCUS", "Buffer_adm"];
--- a/src/HOLCF/IMP/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/HOLCF/IMP/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -4,6 +4,4 @@
     Copyright   1997  TU Muenchen
 *)
 
-with_path "../../HOL/IMP" (no_document time_use_thy) "Natural";
-time_use_thy "HoareEx";
-
+use_thys ["../../HOL/IMP/Natural", "HoareEx"];
--- a/src/HOLCF/ex/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/HOLCF/ex/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -4,11 +4,5 @@
 Misc HOLCF examples.
 *)
 
-time_use_thy "Dnat";
-time_use_thy "Stream";
-time_use_thy "Dagstuhl";
-time_use_thy "Focus_ex";
-time_use_thy "Fix2";
-time_use_thy "Hoare";
-time_use_thy "Loop";
-time_use_thy "Fixrec_ex";
+use_thys ["Dnat", "Stream", "Dagstuhl", "Focus_ex", "Fix2", "Hoare",
+  "Loop", "Fixrec_ex"];
--- a/src/LCF/ex/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/LCF/ex/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -6,7 +6,4 @@
 Some examples from Lawrence Paulson's book Logic and Computation.
 *)
 
-time_use_thy "Ex1";
-time_use_thy "Ex2";
-time_use_thy "Ex3";
-time_use_thy "Ex4";
+use_thys ["Ex1", "Ex2", "Ex3", "Ex4"];
--- a/src/Sequents/LK/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/Sequents/LK/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -6,7 +6,5 @@
 Examples for Classical Logic.
 *)
 
-use_thy "Propositional";
-use_thy "Quantifiers";
-use_thy "Hard_Quantifiers";
-use_thy "Nat";
+use_thys ["Propositional", "Quantifiers", "Hard_Quantifiers", "Nat"];
+
--- a/src/Sequents/ROOT.ML	Tue Jul 31 22:21:22 2007 +0200
+++ b/src/Sequents/ROOT.ML	Tue Jul 31 23:23:28 2007 +0200
@@ -6,19 +6,7 @@
 Classical Sequent Calculus based on Pure Isabelle. 
 *)
 
-val banner = "Sequent Calculii";
-writeln banner;
-
 Unify.trace_bound:= 20;
 Unify.search_bound := 40;
 
-use_thy "LK";
-
-use_thy "ILL";
-use_thy "ILL_predlog";
-use_thy "Washing";
-
-use_thy "Modal0";
-use_thy"T";
-use_thy"S4";
-use_thy"S43";
+use_thys ["LK", "ILL", "ILL_predlog", "Washing", "Modal0", "T", "S4", "S43"];