Regular_Exp.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Fri, 05 Jul 2013 17:19:17 +0100
changeset 379 8c4b6fb43ebe
parent 203 5d724fe0e096
permissions -rw-r--r--
polished more and updated to new isabelle
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     1
(*  Author:     Tobias Nipkow
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     2
    Copyright   1998 TUM
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     3
*)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     4
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     5
header "Regular expressions"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     6
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     7
theory Regular_Exp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     8
imports Regular_Set
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     9
begin
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    10
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    11
datatype 'a rexp =
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    12
  Zero |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    13
  One |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    14
  Atom 'a |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    15
  Plus "('a rexp)" "('a rexp)" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    16
  Times "('a rexp)" "('a rexp)" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    17
  Star "('a rexp)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    18
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
    19
primrec lang :: "'a rexp => 'a lang" where
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    20
"lang Zero = {}" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    21
"lang One = {[]}" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    22
"lang (Atom a) = {[a]}" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    23
"lang (Plus r s) = (lang r) Un (lang s)" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    24
"lang (Times r s) = conc (lang r) (lang s)" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    25
"lang (Star r) = star(lang r)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    26
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    27
primrec atoms :: "'a rexp \<Rightarrow> 'a set" where
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    28
"atoms Zero = {}" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    29
"atoms One = {}" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    30
"atoms (Atom a) = {a}" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    31
"atoms (Plus r s) = atoms r \<union> atoms s" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    32
"atoms (Times r s) = atoms r \<union> atoms s" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    33
"atoms (Star r) = atoms r"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    34
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    35
primrec nullable :: "'a rexp \<Rightarrow> bool" where
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    36
"nullable (Zero) = False" |
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    37
"nullable (One) = True" |
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    38
"nullable (Atom c) = False" |
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    39
"nullable (Plus r1 r2) = (nullable r1 \<or> nullable r2)" |
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    40
"nullable (Times r1 r2) = (nullable r1 \<and> nullable r2)" |
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    41
"nullable (Star r) = True"
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    42
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    43
lemma nullable_iff: "nullable r \<longleftrightarrow> [] \<in> lang r"
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    44
by (induct r) (auto simp add: conc_def split: if_splits)
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 203
diff changeset
    45
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    46
end