--- a/src/HOL/Library/Simps_Case_Conv.thy Fri Sep 06 12:00:58 2013 +0200
+++ b/src/HOL/Library/Simps_Case_Conv.thy Fri Sep 06 12:05:01 2013 +0200
@@ -1,5 +1,5 @@
-(* Title: HOL/Library/Simps_Case_Conv.thy
- Author: Lars Noschinski
+(* Title: HOL/Library/Simps_Case_Conv.thy
+ Author: Lars Noschinski
*)
theory Simps_Case_Conv
--- a/src/HOL/Library/simps_case_conv.ML Fri Sep 06 12:00:58 2013 +0200
+++ b/src/HOL/Library/simps_case_conv.ML Fri Sep 06 12:05:01 2013 +0200
@@ -1,10 +1,10 @@
(* Title: HOL/Library/simps_case_conv.ML
Author: Lars Noschinski, TU Muenchen
- Gerwin Klein, NICTA
+ Author: Gerwin Klein, NICTA
- Converts function specifications between the representation as
- a list of equations (with patterns on the lhs) and a single
- equation (with a nested case expression on the rhs).
+Convert function specifications between the representation as a list
+of equations (with patterns on the lhs) and a single equation (with a
+nested case expression on the rhs).
*)
signature SIMPS_CASE_CONV =