Theory TESL
text‹\chapter[Core TESL: Syntax and Basics]{The Core of the TESL Language: Syntax and Basics}›
theory TESL
imports Main
begin
section‹Syntactic Representation›
text‹
  We define here the syntax of TESL specifications.
›
subsection‹Basic elements of a specification›
text‹
  The following items appear in specifications:
  ▪ Clocks, which are identified by a name.
  ▪ Tag constants are just constants of a type which denotes the metric time space.
›
datatype     clock         = Clk ‹string›
type_synonym instant_index = ‹nat›