| author | wenzelm | 
| Sat, 16 Jul 2011 17:11:49 +0200 | |
| changeset 43845 | d89353d17f54 | 
| parent 43780 | 2cb2310d68b6 | 
| child 45666 | d83797ef0d2d | 
| permissions | -rw-r--r-- | 
| 27968 | 1 | /* Title: Pure/General/position.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Position properties. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 32450 | 10 | object Position | 
| 11 | {
 | |
| 43780 | 12 | type T = Properties.T | 
| 27968 | 13 | |
| 43780 | 14 | val Line = new Properties.Int(Markup.LINE) | 
| 15 | val Offset = new Properties.Int(Markup.OFFSET) | |
| 16 | val End_Offset = new Properties.Int(Markup.END_OFFSET) | |
| 17 | val File = new Properties.String(Markup.FILE) | |
| 18 | val Id = new Properties.Long(Markup.ID) | |
| 31705 | 19 | |
| 38722 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38568diff
changeset | 20 | object Range | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38568diff
changeset | 21 |   {
 | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38568diff
changeset | 22 | def apply(range: Text.Range): T = Offset(range.start) ++ Offset(range.stop) | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38568diff
changeset | 23 | def unapply(pos: T): Option[Text.Range] = | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38568diff
changeset | 24 |       (Offset.unapply(pos), End_Offset.unapply(pos)) match {
 | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38568diff
changeset | 25 | case (Some(start), Some(stop)) if start <= stop => Some(Text.Range(start, stop)) | 
| 39041 | 26 | case (Some(start), None) => Some(Text.Range(start, start + 1)) | 
| 38722 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38568diff
changeset | 27 | case _ => None | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38568diff
changeset | 28 | } | 
| 
ba31936497c2
organized markup properties via apply/unapply patterns;
 wenzelm parents: 
38568diff
changeset | 29 | } | 
| 38872 | 30 | |
| 38887 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 wenzelm parents: 
38872diff
changeset | 31 | object Id_Range | 
| 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 wenzelm parents: 
38872diff
changeset | 32 |   {
 | 
| 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 wenzelm parents: 
38872diff
changeset | 33 | def unapply(pos: T): Option[(Long, Text.Range)] = | 
| 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 wenzelm parents: 
38872diff
changeset | 34 |       (pos, pos) match {
 | 
| 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 wenzelm parents: 
38872diff
changeset | 35 | case (Id(id), Range(range)) => Some((id, range)) | 
| 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 wenzelm parents: 
38872diff
changeset | 36 | case _ => None | 
| 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 wenzelm parents: 
38872diff
changeset | 37 | } | 
| 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 wenzelm parents: 
38872diff
changeset | 38 | } | 
| 
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
 wenzelm parents: 
38872diff
changeset | 39 | |
| 42327 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41483diff
changeset | 40 | private val purge_pos = Map( | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41483diff
changeset | 41 | Markup.DEF_LINE -> Markup.LINE, | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41483diff
changeset | 42 | Markup.DEF_OFFSET -> Markup.OFFSET, | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41483diff
changeset | 43 | Markup.DEF_END_OFFSET -> Markup.END_OFFSET, | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41483diff
changeset | 44 | Markup.DEF_FILE -> Markup.FILE, | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41483diff
changeset | 45 | Markup.DEF_ID -> Markup.ID) | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41483diff
changeset | 46 | |
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41483diff
changeset | 47 | def purge(props: T): T = | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41483diff
changeset | 48 | for ((x, y) <- props if !Markup.POSITION_PROPERTIES(x)) | 
| 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
41483diff
changeset | 49 | yield (if (purge_pos.isDefinedAt(x)) (purge_pos(x), y) else (x, y)) | 
| 27968 | 50 | } |