Matcher.thy
author urbanc
Mon, 30 Jan 2012 09:33:39 +0000
changeset 270 d98435b93063
parent 262 4190df6f4488
permissions -rw-r--r--
slight polishing
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
     1
theory Matcher
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
     2
  imports "Main" 
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
     3
begin
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
     4
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
     5
section {* Sequential Composition of Sets *}
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
     6
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
     7
definition
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
     8
  Seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
     9
where 
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    10
  "A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    11
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    12
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    13
text {* Two Simple Properties about Sequential Composition *}
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    14
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    15
lemma seq_empty [simp]:
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    16
  shows "A ;; {[]} = A"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    17
  and   "{[]} ;; A = A"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    18
by (simp_all add: Seq_def)
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    19
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    20
lemma seq_null [simp]:
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    21
  shows "A ;; {} = {}"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    22
  and   "{} ;; A = {}"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    23
by (simp_all add: Seq_def)
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    24
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    25
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    26
section {* Kleene Star for Sets *}
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    27
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    28
inductive_set
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    29
  Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    30
  for A :: "string set"
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    31
where
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    32
  start[intro]: "[] \<in> A\<star>"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    33
| step[intro]:  "\<lbrakk>s1 \<in> A; s2 \<in> A\<star>\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> A\<star>"
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    34
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    35
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    36
text {* A Standard Property of Star *}
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    37
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    38
lemma star_cases:
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    39
  shows "A\<star> = {[]} \<union> A ;; A\<star>"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    40
unfolding Seq_def
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    41
by (auto) (metis Star.simps)
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    42
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    43
lemma star_decomp: 
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    44
  assumes a: "c # x \<in> A\<star>" 
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    45
  shows "\<exists>a b. x = a @ b \<and> c # a \<in> A \<and> b \<in> A\<star>"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    46
using a
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    47
by (induct x\<equiv>"c # x" rule: Star.induct) 
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    48
   (auto simp add: append_eq_Cons_conv)
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    49
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    50
section {* Left-Quotient of a Set *}
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    51
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    52
definition
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    53
  Der :: "char \<Rightarrow> string set \<Rightarrow> string set"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    54
where
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    55
  "Der c A \<equiv> {s. [c] @ s \<in> A}"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    56
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    57
lemma Der_null [simp]:
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    58
  shows "Der c {} = {}"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    59
unfolding Der_def
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    60
by auto
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    61
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    62
lemma Der_empty [simp]:
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    63
  shows "Der c {[]} = {}"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    64
unfolding Der_def
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    65
by auto
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    66
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    67
lemma Der_char [simp]:
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    68
  shows "Der c {[d]} = (if c = d then {[]} else {})"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    69
unfolding Der_def
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    70
by auto
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    71
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    72
lemma Der_union [simp]:
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    73
  shows "Der c (A \<union> B) = Der c A \<union> Der c B"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    74
unfolding Der_def
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    75
by auto
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    76
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    77
lemma Der_seq [simp]:
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    78
  shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    79
unfolding Der_def Seq_def
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    80
by (auto simp add: Cons_eq_append_conv)
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    81
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    82
lemma Der_star [simp]:
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    83
  shows "Der c (A\<star>) = (Der c A) ;; A\<star>"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    84
proof -    
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    85
  have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    86
    by (simp only: star_cases[symmetric])
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    87
  also have "... = Der c (A ;; A\<star>)"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    88
    by (simp only: Der_union Der_empty) (simp)
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    89
  also have "... = (Der c A) ;; A\<star> \<union> (if [] \<in> A then Der c (A\<star>) else {})"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    90
    by simp
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    91
  also have "... =  (Der c A) ;; A\<star>"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    92
    unfolding Seq_def Der_def
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    93
    by (auto dest: star_decomp)
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    94
  finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    95
qed
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
    96
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    97
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    98
section {* Regular Expressions *}
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
    99
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   100
datatype rexp =
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   101
  NULL
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   102
| EMPTY
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   103
| CHAR char
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   104
| SEQ rexp rexp
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   105
| ALT rexp rexp
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   106
| STAR rexp
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   107
102
5fed809d0fc1 added definition of DERIV and delta
urbanc
parents: 24
diff changeset
   108
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   109
section {* Semantics of Regular Expressions *}
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   110
 
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   111
fun
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   112
  L :: "rexp \<Rightarrow> string set"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   113
where
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   114
  "L (NULL) = {}"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   115
| "L (EMPTY) = {[]}"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   116
| "L (CHAR c) = {[c]}"
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   117
| "L (SEQ r1 r2) = (L r1) ;; (L r2)"
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   118
| "L (ALT r1 r2) = (L r1) \<union> (L r2)"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   119
| "L (STAR r) = (L r)\<star>"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   120
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   121
section {* The Matcher *}
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   122
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   123
fun
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   124
 nullable :: "rexp \<Rightarrow> bool"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   125
where
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   126
  "nullable (NULL) = False"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   127
| "nullable (EMPTY) = True"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   128
| "nullable (CHAR c) = False"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   129
| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   130
| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   131
| "nullable (STAR r) = True"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   132
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   133
fun
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   134
 der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   135
where
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   136
  "der c (NULL) = NULL"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   137
| "der c (EMPTY) = NULL"
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents: 3
diff changeset
   138
| "der c (CHAR c') = (if c = c' then EMPTY else NULL)"
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   139
| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   140
| "der c (SEQ r1 r2) = ALT (SEQ (der c r1) r2) (if nullable r1 then der c r2 else NULL)"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   141
| "der c (STAR r) = SEQ (der c r) (STAR r)"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   142
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   143
fun 
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   144
 derivative :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   145
where
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   146
  "derivative [] r = r"
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents: 3
diff changeset
   147
| "derivative (c # s) r = derivative s (der c r)"
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   148
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   149
fun
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   150
  matcher :: "rexp \<Rightarrow> string \<Rightarrow> bool"
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   151
where
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   152
  "matcher r s = nullable (derivative s r)"
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   153
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   154
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   155
section {* Correctness Proof of the Matcher *}
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   156
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   157
lemma nullable_correctness:
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   158
  shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   159
by (induct r) (auto simp add: Seq_def) 
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   160
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   161
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   162
lemma der_correctness:
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   163
  shows "L (der c r) = Der c (L r)"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   164
by (induct r) 
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   165
   (simp_all add: nullable_correctness)
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   166
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   167
lemma matcher_correctness:
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   168
  shows "matcher r s \<longleftrightarrow> s \<in> L r"
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   169
by (induct s arbitrary: r)
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   170
   (simp_all add: nullable_correctness der_correctness Der_def)
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   171
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   172
section {* Examples *}
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   173
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   174
definition 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   175
  "CHRA \<equiv> CHAR (CHR ''a'')"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   176
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   177
definition 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   178
  "ALT1 \<equiv> ALT CHRA EMPTY"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   179
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   180
definition 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   181
  "SEQ3 \<equiv> SEQ (SEQ ALT1 ALT1) ALT1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   182
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   183
value "matcher SEQ3 ''aaa''"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents: 155
diff changeset
   184
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   185
value "matcher NULL []"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   186
value "matcher (CHAR (CHR ''a'')) [CHR ''a'']"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   187
value "matcher (CHAR a) [a,a]"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   188
value "matcher (STAR (CHAR a)) []"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   189
value "matcher (STAR (CHAR a))  [a,a]"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   190
value "matcher (SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ''abbbbc''"
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   191
value "matcher (SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ''abbcbbc''"
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   192
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   193
section {* Incorrect Matcher - fun-definition rejected *}
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   194
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   195
fun
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   196
  match :: "rexp list \<Rightarrow> string \<Rightarrow> bool"
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   197
where
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   198
  "match [] [] = True"
5
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents: 3
diff changeset
   199
| "match [] (c # s) = False"
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents: 3
diff changeset
   200
| "match (NULL # rs) s = False"  
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents: 3
diff changeset
   201
| "match (EMPTY # rs) s = match rs s"
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents: 3
diff changeset
   202
| "match (CHAR c # rs) [] = False"
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents: 3
diff changeset
   203
| "match (CHAR c # rs) (d # s) = (if c = d then match rs s else False)"         
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents: 3
diff changeset
   204
| "match (ALT r1 r2 # rs) s = (match (r1 # rs) s \<or> match (r2 # rs) s)" 
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents: 3
diff changeset
   205
| "match (SEQ r1 r2 # rs) s = match (r1 # r2 # rs) s"
074d9a4b2bc9 added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
parents: 3
diff changeset
   206
| "match (STAR r # rs) s = (match rs s \<or> match (r # (STAR r) # rs) s)"
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 103
diff changeset
   207
3
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   208
8d657fa3ba2e added simple regexp matcher from Slind et al
urbanc
parents:
diff changeset
   209
end