thys3/PDerivs.thy
changeset 495 f9cdc295ccf7
equal deleted inserted replaced
494:c730d018ebfa 495:f9cdc295ccf7
       
     1    
       
     2 theory PDerivs
       
     3   imports PosixSpec 
       
     4 begin
       
     5 
       
     6 
       
     7 
       
     8 abbreviation
       
     9   "SEQs rs r \<equiv> (\<Union>r' \<in> rs. {SEQ r' r})"
       
    10 
       
    11 lemma SEQs_eq_image:
       
    12   "SEQs rs r = (\<lambda>r'. SEQ r' r) ` rs"
       
    13   by auto
       
    14 
       
    15 fun
       
    16   pder :: "char \<Rightarrow> rexp \<Rightarrow> rexp set"
       
    17 where
       
    18   "pder c ZERO = {}"
       
    19 | "pder c ONE = {}"
       
    20 | "pder c (CH d) = (if c = d then {ONE} else {})"
       
    21 | "pder c (ALT r1 r2) = (pder c r1) \<union> (pder c r2)"
       
    22 | "pder c (SEQ r1 r2) = 
       
    23     (if nullable r1 then SEQs (pder c r1) r2 \<union> pder c r2 else SEQs (pder c r1) r2)"
       
    24 | "pder c (STAR r) = SEQs (pder c r) (STAR r)"
       
    25 
       
    26 fun
       
    27   pders :: "char list \<Rightarrow> rexp \<Rightarrow> rexp set"
       
    28 where
       
    29   "pders [] r = {r}"
       
    30 | "pders (c # s) r = \<Union> (pders s ` pder c r)"
       
    31 
       
    32 abbreviation
       
    33  pder_set :: "char \<Rightarrow> rexp set \<Rightarrow> rexp set"
       
    34 where
       
    35   "pder_set c rs \<equiv> \<Union> (pder c ` rs)"
       
    36 
       
    37 abbreviation
       
    38   pders_set :: "char list \<Rightarrow> rexp set \<Rightarrow> rexp set"
       
    39 where
       
    40   "pders_set s rs \<equiv> \<Union> (pders s ` rs)"
       
    41 
       
    42 lemma pders_append:
       
    43   "pders (s1 @ s2) r = \<Union> (pders s2 ` pders s1 r)"
       
    44 by (induct s1 arbitrary: r) (simp_all)
       
    45 
       
    46 lemma pders_snoc:
       
    47   shows "pders (s @ [c]) r = pder_set c (pders s r)"
       
    48 by (simp add: pders_append)
       
    49 
       
    50 lemma pders_simps [simp]:
       
    51   shows "pders s ZERO = (if s = [] then {ZERO} else {})"
       
    52   and   "pders s ONE = (if s = [] then {ONE} else {})"
       
    53   and   "pders s (ALT r1 r2) = (if s = [] then {ALT r1 r2} else (pders s r1) \<union> (pders s r2))"
       
    54 by (induct s) (simp_all)
       
    55 
       
    56 lemma pders_CHAR:
       
    57   shows "pders s (CH c) \<subseteq> {CH c, ONE}"
       
    58 by (induct s) (simp_all)
       
    59 
       
    60 subsection \<open>Relating left-quotients and partial derivatives\<close>
       
    61 
       
    62 lemma Sequ_UNION_distrib:
       
    63 shows "A ;; \<Union>(M ` I) = \<Union>((\<lambda>i. A ;; M i) ` I)"
       
    64 and   "\<Union>(M ` I) ;; A = \<Union>((\<lambda>i. M i ;; A) ` I)"
       
    65 by (auto simp add: Sequ_def)
       
    66 
       
    67 
       
    68 lemma Der_pder:
       
    69   shows "Der c (L r) = \<Union> (L ` pder c r)"
       
    70 by (induct r) (simp_all add: nullable_correctness Sequ_UNION_distrib)
       
    71   
       
    72 lemma Ders_pders:
       
    73   shows "Ders s (L r) = \<Union> (L ` pders s r)"
       
    74 proof (induct s arbitrary: r)
       
    75   case (Cons c s)
       
    76   have ih: "\<And>r. Ders s (L r) = \<Union> (L ` pders s r)" by fact
       
    77   have "Ders (c # s) (L r) = Ders s (Der c (L r))" by (simp add: Ders_def Der_def)
       
    78   also have "\<dots> = Ders s (\<Union> (L ` pder c r))" by (simp add: Der_pder)
       
    79   also have "\<dots> = (\<Union>A\<in>(L ` (pder c r)). (Ders s A))"
       
    80     by (auto simp add:  Ders_def)
       
    81   also have "\<dots> = \<Union> (L ` (pders_set s (pder c r)))"
       
    82     using ih by auto
       
    83   also have "\<dots> = \<Union> (L ` (pders (c # s) r))" by simp
       
    84   finally show "Ders (c # s) (L r) = \<Union> (L ` pders (c # s) r)" .
       
    85 qed (simp add: Ders_def)
       
    86 
       
    87 subsection \<open>Relating derivatives and partial derivatives\<close>
       
    88 
       
    89 lemma der_pder:
       
    90   shows "\<Union> (L ` (pder c r)) = L (der c r)"
       
    91 unfolding der_correctness Der_pder by simp
       
    92 
       
    93 lemma ders_pders:
       
    94   shows "\<Union> (L ` (pders s r)) = L (ders s r)"
       
    95 unfolding der_correctness ders_correctness Ders_pders by simp
       
    96 
       
    97 
       
    98 subsection \<open>Finiteness property of partial derivatives\<close>
       
    99 
       
   100 definition
       
   101   pders_Set :: "string set \<Rightarrow> rexp \<Rightarrow> rexp set"
       
   102 where
       
   103   "pders_Set A r \<equiv> \<Union>x \<in> A. pders x r"
       
   104 
       
   105 lemma pders_Set_subsetI:
       
   106   assumes "\<And>s. s \<in> A \<Longrightarrow> pders s r \<subseteq> C"
       
   107   shows "pders_Set A r \<subseteq> C"
       
   108 using assms unfolding pders_Set_def by (rule UN_least)
       
   109 
       
   110 lemma pders_Set_union:
       
   111   shows "pders_Set (A \<union> B) r = (pders_Set A r \<union> pders_Set B r)"
       
   112 by (simp add: pders_Set_def)
       
   113 
       
   114 lemma pders_Set_subset:
       
   115   shows "A \<subseteq> B \<Longrightarrow> pders_Set A r \<subseteq> pders_Set B r"
       
   116 by (auto simp add: pders_Set_def)
       
   117 
       
   118 definition
       
   119   "UNIV1 \<equiv> UNIV - {[]}"
       
   120 
       
   121 lemma pders_Set_ZERO [simp]:
       
   122   shows "pders_Set UNIV1 ZERO = {}"
       
   123 unfolding UNIV1_def pders_Set_def by auto
       
   124 
       
   125 lemma pders_Set_ONE [simp]:
       
   126   shows "pders_Set UNIV1 ONE = {}"
       
   127 unfolding UNIV1_def pders_Set_def by (auto split: if_splits)
       
   128 
       
   129 lemma pders_Set_CHAR [simp]:
       
   130   shows "pders_Set UNIV1 (CH c) = {ONE}"
       
   131 unfolding UNIV1_def pders_Set_def
       
   132 apply(auto)
       
   133 apply(frule rev_subsetD)
       
   134 apply(rule pders_CHAR)
       
   135 apply(simp)
       
   136 apply(case_tac xa)
       
   137 apply(auto split: if_splits)
       
   138 done
       
   139 
       
   140 lemma pders_Set_ALT [simp]:
       
   141   shows "pders_Set UNIV1 (ALT r1 r2) = pders_Set UNIV1 r1 \<union> pders_Set UNIV1 r2"
       
   142 unfolding UNIV1_def pders_Set_def by auto
       
   143 
       
   144 
       
   145 text \<open>Non-empty suffixes of a string (needed for the cases of @{const SEQ} and @{const STAR} below)\<close>
       
   146 
       
   147 definition
       
   148   "PSuf s \<equiv> {v. v \<noteq> [] \<and> (\<exists>u. u @ v = s)}"
       
   149 
       
   150 lemma PSuf_snoc:
       
   151   shows "PSuf (s @ [c]) = (PSuf s) ;; {[c]} \<union> {[c]}"
       
   152 unfolding PSuf_def Sequ_def
       
   153 by (auto simp add: append_eq_append_conv2 append_eq_Cons_conv)
       
   154 
       
   155 lemma PSuf_Union:
       
   156   shows "(\<Union>v \<in> PSuf s ;; {[c]}. f v) = (\<Union>v \<in> PSuf s. f (v @ [c]))"
       
   157 by (auto simp add: Sequ_def)
       
   158 
       
   159 lemma pders_Set_snoc:
       
   160   shows "pders_Set (PSuf s ;; {[c]}) r = (pder_set c (pders_Set (PSuf s) r))"
       
   161 unfolding pders_Set_def
       
   162 by (simp add: PSuf_Union pders_snoc)
       
   163 
       
   164 lemma pders_SEQ:
       
   165   shows "pders s (SEQ r1 r2) \<subseteq> SEQs (pders s r1) r2 \<union> (pders_Set (PSuf s) r2)"
       
   166 proof (induct s rule: rev_induct)
       
   167   case (snoc c s)
       
   168   have ih: "pders s (SEQ r1 r2) \<subseteq> SEQs (pders s r1) r2 \<union> (pders_Set (PSuf s) r2)" 
       
   169     by fact
       
   170   have "pders (s @ [c]) (SEQ r1 r2) = pder_set c (pders s (SEQ r1 r2))" 
       
   171     by (simp add: pders_snoc)
       
   172   also have "\<dots> \<subseteq> pder_set c (SEQs (pders s r1) r2 \<union> (pders_Set (PSuf s) r2))"
       
   173     using ih by fastforce
       
   174   also have "\<dots> = pder_set c (SEQs (pders s r1) r2) \<union> pder_set c (pders_Set (PSuf s) r2)"
       
   175     by (simp)
       
   176   also have "\<dots> = pder_set c (SEQs (pders s r1) r2) \<union> pders_Set (PSuf s ;; {[c]}) r2"
       
   177     by (simp add: pders_Set_snoc)
       
   178   also 
       
   179   have "\<dots> \<subseteq> pder_set c (SEQs (pders s r1) r2) \<union> pder c r2 \<union> pders_Set (PSuf s ;; {[c]}) r2"
       
   180     by auto
       
   181   also 
       
   182   have "\<dots> \<subseteq> SEQs (pder_set c (pders s r1)) r2 \<union> pder c r2 \<union> pders_Set (PSuf s ;; {[c]}) r2"
       
   183     by (auto simp add: if_splits)
       
   184   also have "\<dots> = SEQs (pders (s @ [c]) r1) r2 \<union> pder c r2 \<union> pders_Set (PSuf s ;; {[c]}) r2"
       
   185     by (simp add: pders_snoc)
       
   186   also have "\<dots> \<subseteq> SEQs (pders (s @ [c]) r1) r2 \<union> pders_Set (PSuf (s @ [c])) r2"
       
   187     unfolding pders_Set_def by (auto simp add: PSuf_snoc)  
       
   188   finally show ?case .
       
   189 qed (simp) 
       
   190 
       
   191 lemma pders_Set_SEQ_aux1:
       
   192   assumes a: "s \<in> UNIV1"
       
   193   shows "pders_Set (PSuf s) r \<subseteq> pders_Set UNIV1 r"
       
   194 using a unfolding UNIV1_def PSuf_def pders_Set_def by auto
       
   195 
       
   196 lemma pders_Set_SEQ_aux2:
       
   197   assumes a: "s \<in> UNIV1"
       
   198   shows "SEQs (pders s r1) r2 \<subseteq> SEQs (pders_Set UNIV1 r1) r2"
       
   199 using a unfolding pders_Set_def by auto
       
   200 
       
   201 lemma pders_Set_SEQ:
       
   202   shows "pders_Set UNIV1 (SEQ r1 r2) \<subseteq> SEQs (pders_Set UNIV1 r1) r2 \<union> pders_Set UNIV1 r2"
       
   203 apply(rule pders_Set_subsetI)
       
   204 apply(rule subset_trans)
       
   205 apply(rule pders_SEQ)
       
   206 using pders_Set_SEQ_aux1 pders_Set_SEQ_aux2
       
   207 apply auto
       
   208 apply blast
       
   209 done
       
   210 
       
   211 lemma pders_STAR:
       
   212   assumes a: "s \<noteq> []"
       
   213   shows "pders s (STAR r) \<subseteq> SEQs (pders_Set (PSuf s) r) (STAR r)"
       
   214 using a
       
   215 proof (induct s rule: rev_induct)
       
   216   case (snoc c s)
       
   217   have ih: "s \<noteq> [] \<Longrightarrow> pders s (STAR r) \<subseteq> SEQs (pders_Set (PSuf s) r) (STAR r)" by fact
       
   218   { assume asm: "s \<noteq> []"
       
   219     have "pders (s @ [c]) (STAR r) = pder_set c (pders s (STAR r))" by (simp add: pders_snoc)
       
   220     also have "\<dots> \<subseteq> pder_set c (SEQs (pders_Set (PSuf s) r) (STAR r))"
       
   221       using ih[OF asm] by fast
       
   222     also have "\<dots> \<subseteq> SEQs (pder_set c (pders_Set (PSuf s) r)) (STAR r) \<union> pder c (STAR r)"
       
   223       by (auto split: if_splits)
       
   224     also have "\<dots> \<subseteq> SEQs (pders_Set (PSuf (s @ [c])) r) (STAR r) \<union> (SEQs (pder c r) (STAR r))"
       
   225       by (simp only: PSuf_snoc pders_Set_snoc pders_Set_union)
       
   226          (auto simp add: pders_Set_def)
       
   227     also have "\<dots> = SEQs (pders_Set (PSuf (s @ [c])) r) (STAR r)"
       
   228       by (auto simp add: PSuf_snoc PSuf_Union pders_snoc pders_Set_def)
       
   229     finally have ?case .
       
   230   }
       
   231   moreover
       
   232   { assume asm: "s = []"
       
   233     then have ?case by (auto simp add: pders_Set_def pders_snoc PSuf_def)
       
   234   }
       
   235   ultimately show ?case by blast
       
   236 qed (simp)
       
   237 
       
   238 lemma pders_Set_STAR:
       
   239   shows "pders_Set UNIV1 (STAR r) \<subseteq> SEQs (pders_Set UNIV1 r) (STAR r)"
       
   240 apply(rule pders_Set_subsetI)
       
   241 apply(rule subset_trans)
       
   242 apply(rule pders_STAR)
       
   243 apply(simp add: UNIV1_def)
       
   244 apply(simp add: UNIV1_def PSuf_def)
       
   245 apply(auto simp add: pders_Set_def)
       
   246 done
       
   247 
       
   248 lemma finite_SEQs [simp]:
       
   249   assumes a: "finite A"
       
   250   shows "finite (SEQs A r)"
       
   251 using a by auto
       
   252 
       
   253 
       
   254 lemma finite_pders_Set_UNIV1:
       
   255   shows "finite (pders_Set UNIV1 r)"
       
   256 apply(induct r)
       
   257 apply(simp_all add: 
       
   258   finite_subset[OF pders_Set_SEQ]
       
   259   finite_subset[OF pders_Set_STAR])
       
   260 done
       
   261     
       
   262 lemma pders_Set_UNIV:
       
   263   shows "pders_Set UNIV r = pders [] r \<union> pders_Set UNIV1 r"
       
   264 unfolding UNIV1_def pders_Set_def
       
   265 by blast
       
   266 
       
   267 lemma finite_pders_Set_UNIV:
       
   268   shows "finite (pders_Set UNIV r)"
       
   269 unfolding pders_Set_UNIV
       
   270 by (simp add: finite_pders_Set_UNIV1)
       
   271 
       
   272 lemma finite_pders_set:
       
   273   shows "finite (pders_Set A r)"
       
   274 by (metis finite_pders_Set_UNIV pders_Set_subset rev_finite_subset subset_UNIV)
       
   275 
       
   276 
       
   277 text \<open>The following relationship between the alphabetic width of regular expressions
       
   278 (called \<open>awidth\<close> below) and the number of partial derivatives was proved
       
   279 by Antimirov~\cite{Antimirov95} and formalized by Max Haslbeck.\<close>
       
   280 
       
   281 fun awidth :: "rexp \<Rightarrow> nat" where
       
   282 "awidth ZERO = 0" |
       
   283 "awidth ONE = 0" |
       
   284 "awidth (CH a) = 1" |
       
   285 "awidth (ALT r1 r2) = awidth r1 + awidth r2" |
       
   286 "awidth (SEQ r1 r2) = awidth r1 + awidth r2" |
       
   287 "awidth (STAR r1) = awidth r1"
       
   288 
       
   289 lemma card_SEQs_pders_Set_le:
       
   290   shows  "card (SEQs (pders_Set A r) s) \<le> card (pders_Set A r)"
       
   291   using finite_pders_set 
       
   292   unfolding SEQs_eq_image 
       
   293   by (rule card_image_le)
       
   294 
       
   295 lemma card_pders_set_UNIV1_le_awidth: 
       
   296   shows "card (pders_Set UNIV1 r) \<le> awidth r"
       
   297 proof (induction r)
       
   298   case (ALT r1 r2)
       
   299   have "card (pders_Set UNIV1 (ALT r1 r2)) = card (pders_Set UNIV1 r1 \<union> pders_Set UNIV1 r2)" by simp
       
   300   also have "\<dots> \<le> card (pders_Set UNIV1 r1) + card (pders_Set UNIV1 r2)"
       
   301     by(simp add: card_Un_le)
       
   302   also have "\<dots> \<le> awidth (ALT r1 r2)" using ALT.IH by simp
       
   303   finally show ?case .
       
   304 next
       
   305   case (SEQ r1 r2)
       
   306   have "card (pders_Set UNIV1 (SEQ r1 r2)) \<le> card (SEQs (pders_Set UNIV1 r1) r2 \<union> pders_Set UNIV1 r2)"
       
   307     by (simp add: card_mono finite_pders_set pders_Set_SEQ)
       
   308   also have "\<dots> \<le> card (SEQs (pders_Set UNIV1 r1) r2) + card (pders_Set UNIV1 r2)"
       
   309     by (simp add: card_Un_le)
       
   310   also have "\<dots> \<le> card (pders_Set UNIV1 r1) + card (pders_Set UNIV1 r2)"
       
   311     by (simp add: card_SEQs_pders_Set_le)
       
   312   also have "\<dots> \<le> awidth (SEQ r1 r2)" using SEQ.IH by simp
       
   313   finally show ?case .
       
   314 next
       
   315   case (STAR r)
       
   316   have "card (pders_Set UNIV1 (STAR r)) \<le> card (SEQs (pders_Set UNIV1 r) (STAR r))"
       
   317     by (simp add: card_mono finite_pders_set pders_Set_STAR)
       
   318   also have "\<dots> \<le> card (pders_Set UNIV1 r)" by (rule card_SEQs_pders_Set_le)
       
   319   also have "\<dots> \<le> awidth (STAR r)" by (simp add: STAR.IH)
       
   320   finally show ?case .
       
   321 qed (auto)
       
   322 
       
   323 text\<open>Antimirov's Theorem 3.4:\<close>
       
   324 
       
   325 theorem card_pders_set_UNIV_le_awidth: 
       
   326   shows "card (pders_Set UNIV r) \<le> awidth r + 1"
       
   327 proof -
       
   328   have "card (insert r (pders_Set UNIV1 r)) \<le> Suc (card (pders_Set UNIV1 r))"
       
   329     by(auto simp: card_insert_if[OF finite_pders_Set_UNIV1])
       
   330   also have "\<dots> \<le> Suc (awidth r)" by(simp add: card_pders_set_UNIV1_le_awidth)
       
   331   finally show ?thesis by(simp add: pders_Set_UNIV)
       
   332 qed 
       
   333 
       
   334 text\<open>Antimirov's Corollary 3.5:\<close>
       
   335 (*W stands for word set*)
       
   336 corollary card_pders_set_le_awidth: 
       
   337   shows "card (pders_Set W r) \<le> awidth r + 1"
       
   338 proof -
       
   339   have "card (pders_Set W r) \<le> card (pders_Set UNIV r)"
       
   340     by (simp add: card_mono finite_pders_set pders_Set_subset)
       
   341   also have "... \<le> awidth r + 1"
       
   342     by (rule card_pders_set_UNIV_le_awidth)
       
   343   finally show "card (pders_Set W r) \<le> awidth r + 1" by simp
       
   344 qed
       
   345 
       
   346 (* other result by antimirov *)
       
   347 
       
   348 lemma card_pders_awidth: 
       
   349   shows "card (pders s r) \<le> awidth r + 1"
       
   350 proof -
       
   351   have "pders s r \<subseteq> pders_Set UNIV r"
       
   352     using pders_Set_def by auto
       
   353   then have "card (pders s r) \<le> card (pders_Set UNIV r)"
       
   354     by (simp add: card_mono finite_pders_set)
       
   355   then show "card (pders s r) \<le> awidth r + 1"
       
   356     using card_pders_set_le_awidth order_trans by blast
       
   357 qed
       
   358     
       
   359   
       
   360   
       
   361 
       
   362 
       
   363 fun subs :: "rexp \<Rightarrow> rexp set" where
       
   364 "subs ZERO = {ZERO}" |
       
   365 "subs ONE = {ONE}" |
       
   366 "subs (CH a) = {CH a, ONE}" |
       
   367 "subs (ALT r1 r2) = (subs r1 \<union> subs r2 \<union> {ALT r1 r2})" |
       
   368 "subs (SEQ r1 r2) = (subs r1 \<union> subs r2 \<union> {SEQ r1 r2} \<union>  SEQs (subs r1) r2)" |
       
   369 "subs (STAR r1) = (subs r1 \<union> {STAR r1} \<union> SEQs (subs r1) (STAR r1))"
       
   370 
       
   371 lemma subs_finite:
       
   372   shows "finite (subs r)"
       
   373   apply(induct r) 
       
   374   apply(simp_all)
       
   375   done
       
   376 
       
   377 
       
   378 
       
   379 lemma pders_subs:
       
   380   shows "pders s r \<subseteq> subs r"
       
   381   apply(induct r arbitrary: s)
       
   382        apply(simp)
       
   383       apply(simp)
       
   384      apply(simp add: pders_CHAR)
       
   385 (*  SEQ case *)
       
   386     apply(simp)
       
   387     apply(rule subset_trans)
       
   388      apply(rule pders_SEQ)
       
   389     defer
       
   390 (* ALT case *)
       
   391     apply(simp)
       
   392     apply(rule impI)
       
   393     apply(rule conjI)
       
   394   apply blast
       
   395     apply blast
       
   396 (* STAR case *)
       
   397     apply(case_tac s)
       
   398     apply(simp)
       
   399    apply(rule subset_trans)
       
   400   thm pders_STAR
       
   401      apply(rule pders_STAR)
       
   402      apply(simp)
       
   403     apply(auto simp add: pders_Set_def)[1]
       
   404   apply(simp)
       
   405   apply(rule conjI)
       
   406    apply blast
       
   407 apply(auto simp add: pders_Set_def)[1]
       
   408   done
       
   409 
       
   410 fun size2 :: "rexp \<Rightarrow> nat" where
       
   411   "size2 ZERO = 1" |
       
   412   "size2 ONE = 1" |
       
   413   "size2 (CH c) = 1" |
       
   414   "size2 (ALT r1 r2) = Suc (size2 r1 + size2 r2)" |
       
   415   "size2 (SEQ r1 r2) = Suc (size2 r1 + size2 r2)" |
       
   416   "size2 (STAR r1) = Suc (size2 r1)" 
       
   417 
       
   418 
       
   419 lemma size_rexp:
       
   420   fixes r :: rexp
       
   421   shows "1 \<le> size2 r"
       
   422   apply(induct r)
       
   423   apply(simp)
       
   424   apply(simp_all)
       
   425   done
       
   426 
       
   427 lemma subs_size2:
       
   428   shows "\<forall>r1 \<in> subs r. size2 r1 \<le> Suc (size2 r * size2 r)"
       
   429   apply(induct r)
       
   430        apply(simp)
       
   431       apply(simp)
       
   432      apply(simp)
       
   433 (* SEQ case *)
       
   434     apply(simp)
       
   435     apply(auto)[1]
       
   436   apply (smt Suc_n_not_le_n add.commute distrib_left le_Suc_eq left_add_mult_distrib nat_le_linear trans_le_add1)
       
   437   apply (smt Suc_le_mono Suc_n_not_le_n le_trans nat_le_linear power2_eq_square power2_sum semiring_normalization_rules(23) trans_le_add2)
       
   438   apply (smt Groups.add_ac(3) Suc_n_not_le_n distrib_left le_Suc_eq left_add_mult_distrib nat_le_linear trans_le_add1)
       
   439 (*  ALT case  *)
       
   440    apply(simp)
       
   441    apply(auto)[1]
       
   442   apply (smt Groups.add_ac(2) Suc_le_mono Suc_n_not_le_n le_add2 linear order_trans power2_eq_square power2_sum)
       
   443   apply (smt Groups.add_ac(2) Suc_le_mono Suc_n_not_le_n left_add_mult_distrib linear mult.commute order.trans trans_le_add1)
       
   444 (* STAR case *)
       
   445   apply(auto)[1]
       
   446   apply(drule_tac x="r'" in bspec)
       
   447    apply(simp)
       
   448   apply(rule le_trans)
       
   449    apply(assumption)
       
   450   apply(simp)
       
   451   using size_rexp
       
   452   apply(simp)
       
   453   done
       
   454 
       
   455 lemma awidth_size:
       
   456   shows "awidth r \<le> size2 r"
       
   457   apply(induct r)
       
   458        apply(simp_all)
       
   459   done
       
   460 
       
   461 lemma Sum1:
       
   462   fixes A B :: "nat set"
       
   463   assumes "A \<subseteq> B" "finite A" "finite B"
       
   464   shows "\<Sum>A \<le> \<Sum>B"
       
   465   using  assms
       
   466   by (simp add: sum_mono2)
       
   467 
       
   468 lemma Sum2:
       
   469   fixes A :: "rexp set"  
       
   470   and   f g :: "rexp \<Rightarrow> nat" 
       
   471   assumes "finite A" "\<forall>x \<in> A. f x \<le> g x"
       
   472   shows "sum f A \<le> sum g A"
       
   473   using  assms
       
   474   apply(induct A)
       
   475    apply(auto)
       
   476   done
       
   477 
       
   478 
       
   479 
       
   480 
       
   481 
       
   482 lemma pders_max_size:
       
   483   shows "(sum size2 (pders s r)) \<le> (Suc (size2 r)) ^ 3"
       
   484 proof -
       
   485   have "(sum size2 (pders s r)) \<le> sum (\<lambda>_. Suc (size2 r  * size2 r)) (pders s r)" 
       
   486     apply(rule_tac Sum2)
       
   487      apply (meson pders_subs rev_finite_subset subs_finite)
       
   488     using pders_subs subs_size2 by blast
       
   489   also have "... \<le> (Suc (size2 r  * size2 r)) * (sum (\<lambda>_. 1) (pders s r))"
       
   490     by simp
       
   491   also have "... \<le> (Suc (size2 r  * size2 r)) * card (pders s r)"
       
   492     by simp
       
   493   also have "... \<le> (Suc (size2 r  * size2 r)) * (Suc (awidth r))"
       
   494     using Suc_eq_plus1 card_pders_awidth mult_le_mono2 by presburger
       
   495   also have "... \<le> (Suc (size2 r  * size2 r)) * (Suc (size2 r))"
       
   496     using Suc_le_mono awidth_size mult_le_mono2 by presburger
       
   497   also have "... \<le> (Suc (size2 r)) ^ 3"
       
   498     by (smt One_nat_def Suc_1 Suc_mult_le_cancel1 Suc_n_not_le_n antisym_conv le_Suc_eq mult.commute nat_le_linear numeral_3_eq_3 power2_eq_square power2_le_imp_le power_Suc size_rexp)    
       
   499   finally show ?thesis  .
       
   500 qed
       
   501   
       
   502 lemma pders_Set_max_size:
       
   503   shows "(sum size2 (pders_Set A r)) \<le> (Suc (size2 r)) ^ 3"
       
   504 proof -
       
   505   have "(sum size2 (pders_Set A r)) \<le> sum (\<lambda>_. Suc (size2 r  * size2 r)) (pders_Set A r)" 
       
   506     apply(rule_tac Sum2)
       
   507      apply (simp add: finite_pders_set)
       
   508     by (meson pders_Set_subsetI pders_subs subs_size2 subsetD)
       
   509   also have "... \<le> (Suc (size2 r  * size2 r)) * (sum (\<lambda>_. 1) (pders_Set A r))"
       
   510     by simp
       
   511   also have "... \<le> (Suc (size2 r  * size2 r)) * card (pders_Set A r)"
       
   512     by simp
       
   513   also have "... \<le> (Suc (size2 r  * size2 r)) * (Suc (awidth r))"
       
   514     using Suc_eq_plus1 card_pders_set_le_awidth mult_le_mono2 by presburger
       
   515   also have "... \<le> (Suc (size2 r  * size2 r)) * (Suc (size2 r))"
       
   516     using Suc_le_mono awidth_size mult_le_mono2 by presburger
       
   517   also have "... \<le> (Suc (size2 r)) ^ 3"
       
   518     by (smt One_nat_def Suc_1 Suc_mult_le_cancel1 Suc_n_not_le_n antisym_conv le_Suc_eq mult.commute nat_le_linear numeral_3_eq_3 power2_eq_square power2_le_imp_le power_Suc size_rexp)    
       
   519   finally show ?thesis  .
       
   520 qed    
       
   521 
       
   522 fun height :: "rexp \<Rightarrow> nat" where
       
   523   "height ZERO = 1" |
       
   524   "height ONE = 1" |
       
   525   "height (CH c) = 1" |
       
   526   "height (ALT r1 r2) = Suc (max (height r1) (height r2))" |
       
   527   "height (SEQ r1 r2) = Suc (max (height r1) (height r2))" |
       
   528   "height (STAR r1) = Suc (height r1)" 
       
   529 
       
   530 lemma height_size2:
       
   531   shows "height r \<le> size2 r"
       
   532   apply(induct r)
       
   533   apply(simp_all)
       
   534   done
       
   535 
       
   536 lemma height_rexp:
       
   537   fixes r :: rexp
       
   538   shows "1 \<le> height r"
       
   539   apply(induct r)
       
   540   apply(simp_all)
       
   541   done
       
   542 
       
   543 lemma subs_height:
       
   544   shows "\<forall>r1 \<in> subs r. height r1 \<le> Suc (height r)"
       
   545   apply(induct r)
       
   546   apply(auto)+
       
   547   done  
       
   548 
       
   549 fun lin_concat :: "(char \<times> rexp) \<Rightarrow> rexp \<Rightarrow> (char \<times> rexp)" (infixl "[.]" 91)
       
   550   where
       
   551 "(c, ZERO) [.] t = (c, ZERO)"
       
   552 | "(c, ONE) [.] t = (c, t)"
       
   553 | "(c, p) [.] t = (c, SEQ p t)"
       
   554 
       
   555 
       
   556 fun circle_concat :: "(char \<times> rexp ) set \<Rightarrow> rexp \<Rightarrow> (char \<times> rexp) set" ( infixl "\<circle>" 90)
       
   557   where
       
   558 "l \<circle> ZERO = {}"
       
   559 | "l \<circle> ONE = l"
       
   560 | "l \<circle> t  = ( (\<lambda>p.  p [.] t) ` l ) "
       
   561 
       
   562 
       
   563 
       
   564 fun linear_form :: "rexp \<Rightarrow>( char \<times> rexp ) set" 
       
   565   where
       
   566   "linear_form ZERO = {}"
       
   567 | "linear_form ONE = {}"
       
   568 | "linear_form (CH c)  = {(c, ONE)}"
       
   569 | "linear_form (ALT r1 r2) = (linear_form) r1 \<union> (linear_form r2)"
       
   570 | "linear_form (SEQ r1 r2) = (if (nullable r1) then (linear_form r1) \<circle> r2 \<union> linear_form r2 else  (linear_form r1) \<circle> r2) "
       
   571 | "linear_form (STAR r ) = (linear_form r) \<circle> (STAR r)"
       
   572 
       
   573 
       
   574 value "linear_form (SEQ (STAR (CH x)) (STAR (ALT (SEQ (CH x) (CH x)) (CH y)  ))  )"
       
   575 
       
   576 
       
   577 value "linear_form  (STAR (ALT (SEQ (CH x) (CH x)) (CH y)  ))  "
       
   578 
       
   579 fun pdero :: "char \<Rightarrow> rexp \<Rightarrow> rexp set"
       
   580   where
       
   581 "pdero c t  = \<Union> ((\<lambda>(d, p). if d = c then {p} else {}) ` (linear_form t) )"
       
   582 
       
   583 fun pderso :: "char list \<Rightarrow> rexp \<Rightarrow> rexp set"
       
   584   where
       
   585   "pderso [] r = {r}"
       
   586 |  "pderso (c # s) r = \<Union> ( pderso s ` (pdero c r) )"
       
   587 
       
   588 lemma pdero_result: 
       
   589   shows "pdero  c (STAR (ALT (CH c) (SEQ (CH c) (CH c)))) =  {SEQ (CH c)(STAR (ALT (CH c) (SEQ (CH c) (CH c)))),(STAR (ALT (CH c) (SEQ (CH c) (CH c))))}"
       
   590   apply(simp)
       
   591   by auto
       
   592 
       
   593 fun concatLen :: "rexp \<Rightarrow> nat" where
       
   594 "concatLen ZERO = 0" |
       
   595 "concatLen ONE = 0" |
       
   596 "concatLen (CH c) = 0" |
       
   597 "concatLen (SEQ v1 v2) = Suc (max (concatLen v1) (concatLen v2))" |
       
   598 " concatLen (ALT v1 v2) =  max (concatLen v1) (concatLen v2)" |
       
   599 " concatLen (STAR v) = Suc (concatLen v)" 
       
   600 
       
   601 
       
   602 
       
   603 end