thys3/src/RegLangs.thy
changeset 563 c92a41d9c4da
parent 496 f493a20feeb3
equal deleted inserted replaced
562:57e33978e55d 563:c92a41d9c4da
    19 lemma Sequ_empty [simp]:
    19 lemma Sequ_empty [simp]:
    20   shows "A ;; {} = {}"
    20   shows "A ;; {} = {}"
    21   and   "{} ;; A = {}"
    21   and   "{} ;; A = {}"
    22   by (simp_all add: Sequ_def)
    22   by (simp_all add: Sequ_def)
    23 
    23 
       
    24 lemma concI[simp,intro]: "u : A \<Longrightarrow> v : B \<Longrightarrow> u@v : A ;; B"
       
    25 by (auto simp add: Sequ_def)
       
    26 
       
    27 lemma concE[elim]: 
       
    28 assumes "w \<in> A ;; B"
       
    29 obtains u v where "u \<in> A" "v \<in> B" "w = u@v"
       
    30 using assms by (auto simp: Sequ_def)
       
    31 
       
    32 lemma concI_if_Nil2: "[] \<in> B \<Longrightarrow> xs : A \<Longrightarrow> xs \<in> A ;; B"
       
    33 by (metis append_Nil2 concI)
       
    34 
       
    35 lemma conc_assoc: "(A ;; B) ;; C = A ;; (B ;; C)"
       
    36 by (auto elim!: concE) (simp only: append_assoc[symmetric] concI)
       
    37 
       
    38 
       
    39 text \<open>Language power operations\<close>
       
    40 
       
    41 overloading lang_pow == "compow :: nat \<Rightarrow> string set \<Rightarrow> string set"
       
    42 begin
       
    43   primrec lang_pow :: "nat \<Rightarrow> string set \<Rightarrow> string set" where
       
    44   "lang_pow 0 A = {[]}" |
       
    45   "lang_pow (Suc n) A = A ;; (lang_pow n A)"
       
    46 end
       
    47 
       
    48 
       
    49 lemma conc_pow_comm:
       
    50   shows "A ;; (A ^^ n) = (A ^^ n) ;; A"
       
    51 by (induct n) (simp_all add: conc_assoc[symmetric])
       
    52 
       
    53 lemma lang_pow_add: "A ^^ (n + m) = (A ^^ n) ;; (A ^^ m)"
       
    54   by (induct n) (auto simp: conc_assoc)
       
    55 
       
    56 lemma lang_empty: 
       
    57   fixes A::"string set"
       
    58   shows "A ^^ 0 = {[]}"
       
    59   by simp
    24 
    60 
    25 section \<open>Semantic Derivative (Left Quotient) of Languages\<close>
    61 section \<open>Semantic Derivative (Left Quotient) of Languages\<close>
    26 
    62 
    27 definition
    63 definition
    28   Der :: "char \<Rightarrow> string set \<Rightarrow> string set"
    64   Der :: "char \<Rightarrow> string set \<Rightarrow> string set"
    86 lemma Star_Der_Sequ: 
   122 lemma Star_Der_Sequ: 
    87   shows "Der c (A\<star>) \<subseteq> (Der c A) ;; A\<star>"
   123   shows "Der c (A\<star>) \<subseteq> (Der c A) ;; A\<star>"
    88 unfolding Der_def Sequ_def
   124 unfolding Der_def Sequ_def
    89 by(auto simp add: Star_decomp)
   125 by(auto simp add: Star_decomp)
    90 
   126 
       
   127 lemma Der_inter[simp]:   "Der a (A \<inter> B) = Der a A \<inter> Der a B"
       
   128   and Der_compl[simp]:   "Der a (-A) = - Der a A"
       
   129   and Der_Union[simp]:   "Der a (Union M) = Union(Der a ` M)"
       
   130   and Der_UN[simp]:      "Der a (UN x:I. S x) = (UN x:I. Der a (S x))"
       
   131 by (auto simp: Der_def)
    91 
   132 
    92 lemma Der_star[simp]:
   133 lemma Der_star[simp]:
    93   shows "Der c (A\<star>) = (Der c A) ;; A\<star>"
   134   shows "Der c (A\<star>) = (Der c A) ;; A\<star>"
    94 proof -    
   135 proof -    
    95   have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)"  
   136   have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)"  
   101   also have "... =  (Der c A) ;; A\<star>"
   142   also have "... =  (Der c A) ;; A\<star>"
   102     using Star_Der_Sequ by auto
   143     using Star_Der_Sequ by auto
   103   finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
   144   finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
   104 qed
   145 qed
   105 
   146 
       
   147 lemma Der_pow[simp]:
       
   148   shows "Der c (A ^^ n) = (if n = 0 then {} else (Der c A) ;; (A ^^ (n - 1)))"
       
   149   apply(induct n arbitrary: A)
       
   150    apply(auto simp add: Cons_eq_append_conv)
       
   151   by (metis Suc_pred concI_if_Nil2 conc_assoc conc_pow_comm lang_pow.simps(2))
       
   152 
       
   153 
   106 lemma Star_concat:
   154 lemma Star_concat:
   107   assumes "\<forall>s \<in> set ss. s \<in> A"  
   155   assumes "\<forall>s \<in> set ss. s \<in> A"  
   108   shows "concat ss \<in> A\<star>"
   156   shows "concat ss \<in> A\<star>"
   109 using assms by (induct ss) (auto)
   157 using assms by (induct ss) (auto)
   110 
   158 
   117   apply(clarify)
   165   apply(clarify)
   118   by (metis append_Nil concat.simps(2) set_ConsD)
   166   by (metis append_Nil concat.simps(2) set_ConsD)
   119 
   167 
   120 
   168 
   121 
   169 
       
   170 
   122 section \<open>Regular Expressions\<close>
   171 section \<open>Regular Expressions\<close>
   123 
   172 
   124 datatype rexp =
   173 datatype rexp =
   125   ZERO
   174   ZERO
   126 | ONE
   175 | ONE
   127 | CH char
   176 | CH char
   128 | SEQ rexp rexp
   177 | SEQ rexp rexp
   129 | ALT rexp rexp
   178 | ALT rexp rexp
   130 | STAR rexp
   179 | STAR rexp
       
   180 | NTIMES rexp nat
   131 
   181 
   132 section \<open>Semantics of Regular Expressions\<close>
   182 section \<open>Semantics of Regular Expressions\<close>
   133  
   183  
   134 fun
   184 fun
   135   L :: "rexp \<Rightarrow> string set"
   185   L :: "rexp \<Rightarrow> string set"
   138 | "L (ONE) = {[]}"
   188 | "L (ONE) = {[]}"
   139 | "L (CH c) = {[c]}"
   189 | "L (CH c) = {[c]}"
   140 | "L (SEQ r1 r2) = (L r1) ;; (L r2)"
   190 | "L (SEQ r1 r2) = (L r1) ;; (L r2)"
   141 | "L (ALT r1 r2) = (L r1) \<union> (L r2)"
   191 | "L (ALT r1 r2) = (L r1) \<union> (L r2)"
   142 | "L (STAR r) = (L r)\<star>"
   192 | "L (STAR r) = (L r)\<star>"
   143 
   193 | "L (NTIMES r n) = (L r) ^^ n"
   144 
   194 
   145 section \<open>Nullable, Derivatives\<close>
   195 section \<open>Nullable, Derivatives\<close>
   146 
   196 
   147 fun
   197 fun
   148  nullable :: "rexp \<Rightarrow> bool"
   198  nullable :: "rexp \<Rightarrow> bool"
   151 | "nullable (ONE) = True"
   201 | "nullable (ONE) = True"
   152 | "nullable (CH c) = False"
   202 | "nullable (CH c) = False"
   153 | "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
   203 | "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
   154 | "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
   204 | "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
   155 | "nullable (STAR r) = True"
   205 | "nullable (STAR r) = True"
   156 
   206 | "nullable (NTIMES r n) = (if n = 0 then True else nullable r)"
   157 
   207 
   158 fun
   208 fun
   159  der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
   209  der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
   160 where
   210 where
   161   "der c (ZERO) = ZERO"
   211   "der c (ZERO) = ZERO"
   165 | "der c (SEQ r1 r2) = 
   215 | "der c (SEQ r1 r2) = 
   166      (if nullable r1
   216      (if nullable r1
   167       then ALT (SEQ (der c r1) r2) (der c r2)
   217       then ALT (SEQ (der c r1) r2) (der c r2)
   168       else SEQ (der c r1) r2)"
   218       else SEQ (der c r1) r2)"
   169 | "der c (STAR r) = SEQ (der c r) (STAR r)"
   219 | "der c (STAR r) = SEQ (der c r) (STAR r)"
       
   220 | "der c (NTIMES r n) = (if n = 0 then ZERO else SEQ (der c r) (NTIMES r (n - 1)))"
       
   221 
   170 
   222 
   171 fun 
   223 fun 
   172  ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
   224  ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
   173 where
   225 where
   174   "ders [] r = r"
   226   "ders [] r = r"
   175 | "ders (c # s) r = ders s (der c r)"
   227 | "ders (c # s) r = ders s (der c r)"
   176 
   228 
   177 
   229 
       
   230 lemma pow_empty_iff:
       
   231   shows "[] \<in> (L r) ^^ n \<longleftrightarrow> (if n = 0 then True else [] \<in> (L r))"
       
   232   by (induct n) (auto simp add: Sequ_def)
       
   233 
   178 lemma nullable_correctness:
   234 lemma nullable_correctness:
   179   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
   235   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
   180 by (induct r) (auto simp add: Sequ_def) 
   236   by (induct r) (auto simp add: Sequ_def pow_empty_iff) 
   181 
   237 
   182 lemma der_correctness:
   238 lemma der_correctness:
   183   shows "L (der c r) = Der c (L r)"
   239   shows "L (der c r) = Der c (L r)"
   184 by (induct r) (simp_all add: nullable_correctness)
   240   apply (induct r) 
       
   241         apply(auto simp add: nullable_correctness Sequ_def)
       
   242   using Der_def apply force
       
   243   using Der_def apply auto[1]
       
   244   apply (smt (verit, ccfv_SIG) Der_def append_eq_Cons_conv mem_Collect_eq)
       
   245   using Der_def apply force
       
   246   using Der_Sequ Sequ_def by auto
   185 
   247 
   186 lemma ders_correctness:
   248 lemma ders_correctness:
   187   shows "L (ders s r) = Ders s (L r)"
   249   shows "L (ders s r) = Ders s (L r)"
   188   by (induct s arbitrary: r)
   250   by (induct s arbitrary: r)
   189      (simp_all add: Ders_def der_correctness Der_def)
   251      (simp_all add: Ders_def der_correctness Der_def)
   195 lemma ders_snoc:
   257 lemma ders_snoc:
   196   shows "ders (s @ [c]) r = der c (ders s r)"
   258   shows "ders (s @ [c]) r = der c (ders s r)"
   197   by (simp add: ders_append)
   259   by (simp add: ders_append)
   198 
   260 
   199 
   261 
   200 (*
       
   201 datatype ctxt = 
       
   202     SeqC rexp bool
       
   203   | AltCL rexp
       
   204   | AltCH rexp 
       
   205   | StarC rexp 
       
   206 
       
   207 function
       
   208      down :: "char \<Rightarrow> rexp \<Rightarrow> ctxt list \<Rightarrow> rexp * ctxt list"
       
   209 and  up :: "char \<Rightarrow> rexp \<Rightarrow> ctxt list \<Rightarrow> rexp * ctxt list"
       
   210 where
       
   211   "down c (SEQ r1 r2) ctxts =
       
   212      (if (nullable r1) then down c r1 (SeqC r2 True # ctxts) 
       
   213       else down c r1 (SeqC r2 False # ctxts))"
       
   214 | "down c (CH d) ctxts = 
       
   215      (if c = d then up c ONE ctxts else up c ZERO ctxts)"
       
   216 | "down c ONE ctxts = up c ZERO ctxts"
       
   217 | "down c ZERO ctxts = up c ZERO ctxts"
       
   218 | "down c (ALT r1 r2) ctxts = down c r1 (AltCH r2 # ctxts)"
       
   219 | "down c (STAR r1) ctxts = down c r1 (StarC r1 # ctxts)"
       
   220 | "up c r [] = (r, [])"
       
   221 | "up c r (SeqC r2 False # ctxts) = up c (SEQ r r2) ctxts"
       
   222 | "up c r (SeqC r2 True # ctxts) = down c r2 (AltCL (SEQ r r2) # ctxts)"
       
   223 | "up c r (AltCL r1 # ctxts) = up c (ALT r1 r) ctxts"
       
   224 | "up c r (AltCH r2 # ctxts) = down c r2 (AltCL r # ctxts)"
       
   225 | "up c r (StarC r1 # ctxts) = up c (SEQ r (STAR r1)) ctxts"
       
   226   apply(pat_completeness)
       
   227   apply(auto)
       
   228   done
       
   229 
       
   230 termination
       
   231   sorry
       
   232 
       
   233 *)
       
   234 
       
   235 
       
   236 end
   262 end