eliminated odd dislocation of keyword declaration and implementation (leftover from 318cd8ac1817)
authorhaftmann
Sun, 16 Feb 2014 13:56:48 +0100
changeset 55509 bd67ebe275e0
parent 55508 90c42b130652
child 55517 a3870c12f254
eliminated odd dislocation of keyword declaration and implementation (leftover from 318cd8ac1817)
src/HOL/Groebner_Basis.thy
src/HOL/Metis.thy
--- a/src/HOL/Groebner_Basis.thy	Sat Feb 15 21:11:29 2014 +0100
+++ b/src/HOL/Groebner_Basis.thy	Sun Feb 16 13:56:48 2014 +0100
@@ -6,6 +6,7 @@
 
 theory Groebner_Basis
 imports Semiring_Normalization
+keywords "try0" :: diag
 begin
 
 subsection {* Groebner Bases *}
--- a/src/HOL/Metis.thy	Sat Feb 15 21:11:29 2014 +0100
+++ b/src/HOL/Metis.thy	Sun Feb 16 13:56:48 2014 +0100
@@ -8,7 +8,6 @@
 
 theory Metis
 imports ATP
-keywords "try0" :: diag
 begin
 
 ML_file "~~/src/Tools/Metis/metis.ML"